Skip to content

Commit 4b3966d

Browse files
committed
observe, fixed an error in _≅e_
1 parent ad0450c commit 4b3966d

File tree

5 files changed

+62
-7
lines changed

5 files changed

+62
-7
lines changed

Core.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -207,7 +207,7 @@ var i₁ , x₁ ≅s var i₂ , x₂ = x₁ ≅ x₂
207207
_ ≅s _ = bot
208208

209209
var i₁ , q₁ ≅e var i₂ , q₂ = i₁ ≅ i₂
210-
π A₁ D₁ , x₁ , e₁ ≅e π A₂ D₂ , x₂ , e₂ = D₁ x₁ , e₁ ≅e D₂ x₂ , e₂
210+
π A₁ D₁ , x₁ , e₁ ≅e π A₂ D₂ , x₂ , e₂ = x₁ ≅ x₂ & D₁ x₁ , e₁ ≅e D₂ x₂ , e₂
211211
(D₁ ⊛ E₁) , s₁ , e₁ ≅e (D₂ ⊛ E₂) , s₂ , e₂ = D₁ , s₁ ≅s D₂ , s₂ & E₁ , e₁ ≅e E₂ , e₂
212212
_ ≅e _ = bot
213213

Data/Maybe.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,6 @@
11
module OTT.Data.Maybe where
22

3-
open import OTT.Main
3+
open import OTT.Main hiding (Maybe; nothing; just)
44

55
maybe : {a} {α : Level a} -> Univ α -> Type₋₁ α
66
maybe {α = α} A = cmu₋₁ α

Main.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,6 @@
11
module OTT.Main where
22

3-
open import OTT.Core hiding (Maybe; nothing; just) public
3+
open import OTT.Core public
44
open import OTT.Coerce public
55
open import OTT.Data.List public
66
open import OTT.Function.Pi public

Prelude.agda

+6-2
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,8 @@ module OTT.Prelude where
33
open import Level
44
renaming (Level to MetaLevel; zero to lzeroₘ; suc to lsucₘ; _⊔_ to _⊔ₘ_) using () public
55
open import Function public
6-
open import Relation.Binary.PropositionalEquality
7-
as P renaming (refl to prefl; trans to ptrans; cong to pcong; cong₂ to pcong₂) using (_≡_) public
6+
open import Relation.Binary.PropositionalEquality as P using (_≡_)
7+
renaming (refl to prefl; trans to ptrans; subst to psubst; cong to pcong; cong₂ to pcong₂) public
88
open import Data.Empty public
99
open import Data.Nat.Base hiding (_⊔_; _≟_; erase) public
1010
open import Data.Maybe.Base using (Maybe; nothing; just) public
@@ -36,6 +36,10 @@ instance
3636
pright : {α} {A : Set α} {x y z : A} -> x ≡ y -> x ≡ z -> y ≡ z
3737
pright prefl prefl = prefl
3838

39+
hpcong₂ : {α β γ} {A : Set α} {B : A -> Set β} {C : Set γ} {x₁ x₂} {y₁ : B x₁} {y₂ : B x₂}
40+
-> (f : x -> B x -> C) -> (q : x₁ ≡ x₂) -> psubst B q y₁ ≡ y₂ -> f x₁ y₁ ≡ f x₂ y₂
41+
hpcong₂ f prefl prefl = prefl
42+
3943
record Apply {α β} {A : Set α} (B : A -> Set β) x : Set β where
4044
constructor tag
4145
field detag : B x

Property/Eq.agda

+53-2
Original file line numberDiff line numberDiff line change
@@ -14,8 +14,6 @@ module _ where
1414
contr : {A : Prop} {x y : ⟦ A ⟧} -> x ≡ y
1515
contr = trustMe
1616

17-
-- We could compare functions with a finite domain for equality,
18-
-- but then equality can't be `_≡_`.
1917
SemEq : {i a} {ι : Level i} {α : Level a} {I : Type ι} -> Desc I α -> Set
2018
SemEq (var i) =
2119
SemEq (π A D) =
@@ -71,6 +69,59 @@ mutual
7169
_≟_ {A = desc I α} {{()}}
7270
_≟_ {A = imu D j } d₁ d₂ = decMu d₁ d₂
7371

72+
coerceFamEnum : {n} {e f : Apply Enum n} -> (A : Apply Enum n -> Set) -> ⟦ e ≅ f ⟧ -> A e -> A f
73+
coerceFamEnum {0} {tag () } {tag () } A q x
74+
coerceFamEnum {1} {tag tt } {tag tt } A q x = x
75+
coerceFamEnum {suc (suc n)} {tag nothing } {tag nothing } A q x = x
76+
coerceFamEnum {suc (suc n)} {tag (just _)} {tag (just _)} A q x =
77+
coerceFamEnum (A ∘ tag ∘ just ∘ detag) q x
78+
coerceFamEnum {suc (suc n)} {tag nothing } {tag (just _)} A () x
79+
coerceFamEnum {suc (suc n)} {tag (just _)} {tag nothing } A () x
80+
81+
mutual
82+
observeSem : {i a} {ι : Level i} {α : Level a}
83+
{I : Type ι} {E : Desc I α} {{eqE : ExtendEq E}}
84+
-> (D : Desc I α) {{edD : SemEq D}}
85+
-> (d₁ d₂ : ⟦ D ⟧ᵈ (μ E)) -> ⟦ D , d₁ ≅s D , d₂ ⟧ -> d₁ ≡ d₂
86+
observeSem (var i) d₁ d₂ q = observe q
87+
observeSem (π A D) {{()}} d₁ d₂ q
88+
observeSem (D ⊛ E) {{eqD , eqE}} (d₁ , e₁) (d₂ , e₂) (qd , qe) =
89+
pcong₂ _,_ (observeSem D {{eqD}} d₁ d₂ qd) (observeSem E {{eqE}} e₁ e₂ qe)
90+
91+
observeExtend : {i a} {ι : Level i} {α : Level a} {I : Type ι}
92+
{E : Desc I α} {j} {{edE : ExtendEq E}}
93+
-> (D : Desc I α) {{edD : ExtendEq D}}
94+
-> (e₁ e₂ : Extend D (μ E) j) -> ⟦ D , e₁ ≅e D , e₂ ⟧ -> e₁ ≡ e₂
95+
observeExtend (var i) q₁ q₂ q = contr
96+
observeExtend (π A D) {{eqA , eqD}} (x₁ , e₁) (x₂ , e₂) (qx , qe)
97+
rewrite observe {x = x₁} {x₂} {{eqA}} qx =
98+
pcong (_,_ _) (observeExtend (D x₂) {{apply eqD x₂}} e₁ e₂ qe)
99+
observeExtend (D ⊛ E) {{eqD , eqE}} (d₁ , e₁) (d₂ , e₂) (qd , qe) =
100+
pcong₂ _,_ (observeSem D {{eqD}} d₁ d₂ qd) (observeExtend E {{eqE}} e₁ e₂ qe)
101+
102+
observeMu : {i a} {ι : Level i} {α : Level a} {I : Type ι} {j}
103+
{D : Desc I α} {d e : μ D j} {{eqD : ExtendEq D}}
104+
-> ⟦ d ≅ e ⟧ -> d ≡ e
105+
observeMu {D = D} {node e₁} {node e₂} q = pcong node (observeExtend D e₁ e₂ q)
106+
107+
observe : {a} {α : Level a} {A : Univ α} {x y : ⟦ A ⟧} {{eqA : Eq A}} -> ⟦ x ≅ y ⟧ -> x ≡ y
108+
observe {A = bot } {()} {()}
109+
observe {A = top } q = prefl
110+
observe {A = nat } q = coerceFamℕ (_ ≡_) q prefl
111+
observe {A = enum n } {e₁} {e₂} q = coerceFamEnum (_ ≡_) q prefl
112+
observe {A = univ α } {{()}}
113+
observe {A = σ A B } {x₁ , y₁} {x₂ , y₂} {{eqA , eqB}} (qx , qy)
114+
rewrite observe {x = x₁} {x₂} {{eqA}} qx = pcong (_,_ _) (observe {{apply eqB x₂}} qy)
115+
observe {A = π A B } {{()}}
116+
observe {A = desc I α} {{()}}
117+
observe {A = imu D j } {node e₁} {node e₂} q = observeMu q
118+
119+
module _ where
120+
open import Relation.Binary.PropositionalEquality.TrustMe
121+
122+
eobserve : {a} {α : Level a} {A : Univ α} {x y : ⟦ A ⟧} {{eqA : Eq A}} -> ⟦ x ≅ y ⟧ -> x ≡ y
123+
eobserve = erase ∘ observe
124+
74125
private
75126
module Test where
76127
open import OTT.Data.Fin

0 commit comments

Comments
 (0)