Skip to content

Commit 2e1ee97

Browse files
committed
Add terminal object of RT
1 parent 8f196fc commit 2e1ee97

File tree

2 files changed

+129
-0
lines changed

2 files changed

+129
-0
lines changed

Diff for: src/Realizability/Topos/FunctionalRelation.agda

+1
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,4 @@
1+
{-# OPTIONS --allow-unsolved-metas #-}
12
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
23
open import Realizability.CombinatoryAlgebra
34
open import Cubical.Foundations.Prelude

Diff for: src/Realizability/Topos/TerminalObject.agda

+128
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,128 @@
1+
open import Realizability.ApplicativeStructure renaming (Term to ApplStrTerm; λ*-naturality to `λ*ComputationRule; λ*-chain to `λ*) hiding (λ*)
2+
open import Realizability.CombinatoryAlgebra
3+
open import Cubical.Foundations.Prelude
4+
open import Cubical.Foundations.HLevels
5+
open import Cubical.Data.Unit
6+
open import Cubical.Data.Empty
7+
open import Cubical.Data.Fin
8+
open import Cubical.Data.Vec
9+
open import Cubical.HITs.PropositionalTruncation
10+
open import Cubical.HITs.PropositionalTruncation.Monad
11+
open import Cubical.HITs.SetQuotients renaming (elimProp to setQuotElimProp; elimProp2 to setQuotElimProp2)
12+
open import Cubical.Categories.Category
13+
open import Cubical.Categories.Limits.Terminal
14+
15+
module Realizability.Topos.TerminalObject
16+
{ℓ ℓ' ℓ''}
17+
{A : Type ℓ}
18+
(ca : CombinatoryAlgebra A)
19+
(isNonTrivial : CombinatoryAlgebra.s ca ≡ CombinatoryAlgebra.k ca ⊥) where
20+
21+
open CombinatoryAlgebra ca
22+
open import Realizability.Tripos.Prealgebra.Predicate.Base ca renaming (Predicate to Predicate')
23+
open import Realizability.Tripos.Prealgebra.Predicate.Properties ca
24+
open import Realizability.Topos.Object {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
25+
open import Realizability.Topos.FunctionalRelation {ℓ' = ℓ'} {ℓ'' = ℓ''} ca isNonTrivial
26+
27+
open Combinators ca renaming (i to Id; ia≡a to Ida≡a)
28+
open PartialEquivalenceRelation
29+
open Predicate' renaming (isSetX to isSetPredicateBase)
30+
private
31+
Predicate = Predicate' {ℓ' = ℓ'} {ℓ'' = ℓ''}
32+
λ*ComputationRule = `λ*ComputationRule as fefermanStructure
33+
λ* = `λ* as fefermanStructure
34+
35+
terminalPartialEquivalenceRelation : PartialEquivalenceRelation Unit*
36+
isSetX terminalPartialEquivalenceRelation = isSetUnit*
37+
isSetPredicateBase (equality terminalPartialEquivalenceRelation) = isSet× isSetUnit* isSetUnit*
38+
∣ equality terminalPartialEquivalenceRelation ∣ (tt* , tt*) r = Unit*
39+
isPropValued (equality terminalPartialEquivalenceRelation) (tt* , tt*) r = isPropUnit*
40+
isSymmetric terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* _ tt* tt* }))
41+
isTransitive terminalPartialEquivalenceRelation = return (k , (λ { tt* tt* tt* _ _ tt* tt* tt* }))
42+
43+
open FunctionalRelation
44+
-- I have officially taken the inlining too far
45+
-- TODO : Refactor
46+
isTerminalTerminalPartialEquivalenceRelation : {Y : Type ℓ'} (perY : PartialEquivalenceRelation Y) isContr (RTMorphism perY terminalPartialEquivalenceRelation)
47+
isTerminalTerminalPartialEquivalenceRelation {Y} perY =
48+
inhProp→isContr
49+
[ record
50+
{ relation =
51+
record
52+
{ isSetX = isSet× (perY .isSetX) isSetUnit*
53+
; ∣_∣ = λ { (y , tt*) r r ⊩ ∣ perY .equality ∣ (y , y) }
54+
; isPropValued = λ { (y , tt*) r perY .equality .isPropValued _ _ } }
55+
; isStrict =
56+
let
57+
prover : ApplStrTerm as 1
58+
prover = ` pair ̇ # fzero ̇ # fzero
59+
in
60+
return
61+
((λ* prover) ,
62+
(λ { y tt* r r⊩y~y
63+
subst
64+
(λ r' r' ⊩ ∣ perY .equality ∣ (y , y))
65+
(sym
66+
(pr₁ ⨾ (λ* prover ⨾ r)
67+
≡⟨ cong (λ x pr₁ ⨾ x) (λ*ComputationRule prover (r ∷ [])) ⟩
68+
pr₁ ⨾ (pair ⨾ r ⨾ r)
69+
≡⟨ pr₁pxy≡x _ _ ⟩
70+
r
71+
∎))
72+
r⊩y~y ,
73+
tt* }))
74+
; isRelational =
75+
do
76+
(trY , trY⊩isTransitiveY) perY .isTransitive
77+
(smY , smY⊩isSymmetricY) perY .isSymmetric
78+
let
79+
prover : ApplStrTerm as 1
80+
prover = ` trY ̇ (` pair ̇ (` smY ̇ (` pr₁ ̇ # fzero)) ̇ (` pr₁ ̇ # fzero))
81+
return
82+
(λ* prover ,
83+
(λ { y y' tt* tt* a b c a⊩y~y' b⊩y~y tt*
84+
let
85+
proofEq : λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ≡ trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a)
86+
proofEq =
87+
λ* prover ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c))
88+
≡⟨ λ*ComputationRule prover ((pair ⨾ a ⨾ (pair ⨾ b ⨾ c)) ∷ []) ⟩
89+
(trY ⨾ (pair ⨾ (smY ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))) ⨾ (pr₁ ⨾ (pair ⨾ a ⨾ (pair ⨾ b ⨾ c)))))
90+
≡⟨ cong₂ (λ x y trY ⨾ (pair ⨾ (smY ⨾ x) ⨾ y)) (pr₁pxy≡x _ _) (pr₁pxy≡x _ _) ⟩
91+
trY ⨾ (pair ⨾ (smY ⨾ a) ⨾ a)
92+
93+
in
94+
subst
95+
(λ r r ⊩ ∣ perY .equality ∣ (y' , y'))
96+
(sym proofEq)
97+
(trY⊩isTransitiveY y' y y' (smY ⨾ a) a (smY⊩isSymmetricY y y' a a⊩y~y') a⊩y~y') }))
98+
; isSingleValued = return (k , (λ { _ tt* tt* _ _ _ _ tt* })) -- nice
99+
; isTotal = return (Id , (λ y r r⊩y~y return (tt* , subst (λ r r ⊩ ∣ perY .equality ∣ (y , y)) (sym (Ida≡a _)) r⊩y~y)))
100+
} ]
101+
λ f g
102+
setQuotElimProp2
103+
(λ f g squash/ f g)
104+
(λ F G
105+
eq/
106+
F G
107+
let
108+
F≤G : pointwiseEntailment perY terminalPartialEquivalenceRelation F G
109+
F≤G =
110+
(do
111+
(tlG , tlG⊩isTotalG) G .isTotal
112+
(stF , stF⊩isStrictF) F .isStrict
113+
let
114+
prover : ApplStrTerm as 1
115+
prover = ` tlG ̇ (` pr₁ ̇ (` stF ̇ # fzero))
116+
return
117+
(λ* prover ,
118+
(λ { y tt* r r⊩Fx
119+
transport
120+
(propTruncIdempotent (G .relation .isPropValued _ _))
121+
(tlG⊩isTotalG y (pr₁ ⨾ (stF ⨾ r)) (stF⊩isStrictF y tt* r r⊩Fx .fst)
122+
>>= λ { (tt* , ⊩Gy) return (subst (λ r' r' ⊩ ∣ G .relation ∣ (y , tt*)) (sym (λ*ComputationRule prover (r ∷ []))) ⊩Gy) }) })))
123+
in F≤G , (F≤G→G≤F perY terminalPartialEquivalenceRelation F G F≤G))
124+
f g
125+
126+
TerminalRT : Terminal RT
127+
TerminalRT =
128+
(Unit* , terminalPartialEquivalenceRelation) , (λ { (Y , perY) isTerminalTerminalPartialEquivalenceRelation perY})

0 commit comments

Comments
 (0)