-
Notifications
You must be signed in to change notification settings - Fork 246
/
Copy pathStrictPartialOrder.agda
63 lines (55 loc) · 1.61 KB
/
StrictPartialOrder.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
------------------------------------------------------------------------
-- The Agda standard library
--
-- Convenient syntax for "equational reasoning" using a strict partial
-- order.
------------------------------------------------------------------------
-- Example uses:
--
-- u≤x : u ≤ x
-- u≤x = begin
-- u ≈⟨ u≈v ⟩
-- v ≡⟨ v≡w ⟩
-- w <⟨ w≤x ⟩
-- x ∎
--
-- u<x : u < x
-- u<x = begin-strict
-- u ≈⟨ u≈v ⟩
-- v ≡⟨ v≡w ⟩
-- w <⟨ w≤x ⟩
-- x ∎
--
-- u<e : u < e
-- u<e = begin-strict
-- u ≈⟨ u≈v ⟩
-- v ≡⟨ v≡w ⟩
-- w <⟨ w<x ⟩
-- x ≤⟨ x≤y ⟩
-- y <⟨ y<z ⟩
-- z ≡˘⟨ d≡z ⟩
-- d ≈˘⟨ e≈d ⟩
-- e ∎
--
-- u≈w : u ≈ w
-- u≈w = begin-equality
-- u ≈⟨ u≈v ⟩
-- v ≡⟨ v≡w ⟩
-- w ∎
{-# OPTIONS --cubical-compatible --safe #-}
open import Relation.Binary.Bundles using (StrictPartialOrder)
module Relation.Binary.Reasoning.StrictPartialOrder
{p₁ p₂ p₃} (S : StrictPartialOrder p₁ p₂ p₃) where
open StrictPartialOrder S
import Relation.Binary.Construct.StrictToNonStrict _≈_ _<_ as NonStrict
------------------------------------------------------------------------
-- Publicly re-export the contents of the base module
open import Relation.Binary.Reasoning.Base.Triple
(NonStrict.isPreorder₂ isStrictPartialOrder)
irrefl
trans
<-resp-≈
NonStrict.<⇒≤
(NonStrict.<-≤-trans trans <-respʳ-≈)
(NonStrict.≤-<-trans Eq.sym trans <-respˡ-≈)
public