Skip to content

Commit 84c09d3

Browse files
authored
Merge pull request #523 from LuuBluum/diffInt
DiffInt≡Int
2 parents 7affe55 + 9fe5fb6 commit 84c09d3

File tree

1 file changed

+71
-4
lines changed

1 file changed

+71
-4
lines changed

Cubical/Data/DiffInt/Base.agda

+71-4
Original file line numberDiff line numberDiff line change
@@ -3,15 +3,82 @@ module Cubical.Data.DiffInt.Base where
33

44
open import Cubical.Foundations.Prelude
55

6-
open import Cubical.HITs.SetQuotients.Base
6+
open import Cubical.HITs.SetQuotients
7+
open import Cubical.Foundations.Isomorphism
78

89
open import Cubical.Data.Sigma
9-
open import Cubical.Data.Nat
10+
open import Cubical.Data.Nat hiding (+-comm ; +-assoc) renaming (_+_ to _+ℕ_)
11+
open import Cubical.Data.Int
1012

1113
rel : (ℕ × ℕ) (ℕ × ℕ) Type₀
1214
rel (a₀ , b₀) (a₁ , b₁) = x ≡ y
1315
where
14-
x = a₀ + b₁
15-
y = a₁ + b₀
16+
x = a₀ + b₁
17+
y = a₁ + b₀
1618

1719
= (ℕ × ℕ) / rel
20+
21+
22+
-- Proof of equivalence between Int and DiffInt
23+
24+
private
25+
-- Prove all the identities for ℕ×ℕ first and use that to build to ℤ
26+
27+
Int→ℕ×ℕ : Int (ℕ × ℕ)
28+
Int→ℕ×ℕ (pos n) = ( n , zero )
29+
Int→ℕ×ℕ (negsuc n) = ( zero , suc n )
30+
31+
ℕ×ℕ→Int : (ℕ × ℕ) Int
32+
ℕ×ℕ→Int(n , m) = n ℕ- m
33+
34+
relInt→ℕ×ℕ : m n rel (Int→ℕ×ℕ (m ℕ- n)) (m , n)
35+
relInt→ℕ×ℕ zero zero = refl
36+
relInt→ℕ×ℕ zero (suc n) = refl
37+
relInt→ℕ×ℕ (suc m) zero = refl
38+
relInt→ℕ×ℕ (suc m) (suc n) =
39+
fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ suc n ≡⟨ +-suc (fst (Int→ℕ×ℕ (m ℕ- n))) n ⟩
40+
suc (fst (Int→ℕ×ℕ (m ℕ- n)) +ℕ n) ≡⟨ cong suc (relInt→ℕ×ℕ m n) ⟩
41+
suc (m +ℕ snd (Int→ℕ×ℕ (m ℕ- n))) ∎
42+
43+
Int→ℕ×ℕ→Int : n ℕ×ℕ→Int (Int→ℕ×ℕ n) ≡ n
44+
Int→ℕ×ℕ→Int (pos n) = refl
45+
Int→ℕ×ℕ→Int (negsuc n) = refl
46+
47+
ℕ×ℕ→Int→ℕ×ℕ : p rel (Int→ℕ×ℕ (ℕ×ℕ→Int p)) p
48+
ℕ×ℕ→Int→ℕ×ℕ (p₀ , p₁) = relInt→ℕ×ℕ p₀ p₁
49+
50+
-- Now build to ℤ using the above
51+
52+
Int→ℤ : Int
53+
Int→ℤ n = [ Int→ℕ×ℕ n ]
54+
55+
ℤ→Int : Int
56+
ℤ→Int [ z ] = ℕ×ℕ→Int z
57+
ℤ→Int(eq/ a b r i) = lemℤeq a b r i
58+
where lemℤeq : (a b : (ℕ × ℕ)) rel a b ℕ×ℕ→Int(a) ≡ ℕ×ℕ→Int(b)
59+
lemℤeq (a₀ , a₁) (b₀ , b₁) r = a₀ ℕ- a₁ ≡⟨ pos- a₀ a₁ ⟩
60+
pos a₀ - pos a₁ ≡[ i ]⟨ ((pos a₀ - pos a₁) + -Cancel (pos b₁) (~ i)) ⟩
61+
(pos a₀ - pos a₁) + (pos b₁ - pos b₁) ≡⟨ +-assoc (pos a₀ + (- pos a₁)) (pos b₁) (- pos b₁) ⟩
62+
((pos a₀ - pos a₁) + pos b₁) - pos b₁ ≡[ i ]⟨ +-assoc (pos a₀) (- pos a₁) (pos b₁) (~ i) + (- pos b₁) ⟩
63+
(pos a₀ + ((- pos a₁) + pos b₁)) - pos b₁ ≡[ i ]⟨ (pos a₀ + +-comm (- pos a₁) (pos b₁) i) - pos b₁ ⟩
64+
(pos a₀ + (pos b₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ +-assoc (pos a₀) (pos b₁) (- pos a₁) i + (- pos b₁) ⟩
65+
((pos a₀ + pos b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ a₀ b₁ (~ i) - pos a₁) - pos b₁ ⟩
66+
(pos (a₀ +ℕ b₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos (r i) - pos a₁) - pos b₁ ⟩
67+
(pos (b₀ +ℕ a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ (pos+ b₀ a₁ i - pos a₁) - pos b₁ ⟩
68+
((pos b₀ + pos a₁) - pos a₁) - pos b₁ ≡[ i ]⟨ +-assoc (pos b₀) (pos a₁) (- pos a₁) (~ i) + (- pos b₁) ⟩
69+
(pos b₀ + (pos a₁ - pos a₁)) - pos b₁ ≡[ i ]⟨ (pos b₀ + (-Cancel (pos a₁) i)) - pos b₁ ⟩
70+
pos b₀ - pos b₁ ≡[ i ]⟨ pos- b₀ b₁ (~ i) ⟩
71+
b₀ ℕ- b₁ ∎
72+
ℤ→Int(squash/ x x₀ p q i j) = isSetInt (ℤ→Int x) (ℤ→Int x₀) (cong ℤ→Int p) (cong ℤ→Int q) i j
73+
74+
Int→ℤ→Int : z ℤ→Int (Int→ℤ z) ≡ z
75+
Int→ℤ→Int (pos n) = refl
76+
Int→ℤ→Int (negsuc n) = refl
77+
78+
ℤ→Int→ℤ : z Int→ℤ (ℤ→Int z) ≡ z
79+
ℤ→Int→ℤ = elimProp (λ z squash/ (Int→ℤ (ℤ→Int z)) z) ℕ×ℕprf
80+
where ℕ×ℕprf : (a : ℕ × ℕ) Int→ℤ (ℤ→Int [ a ]) ≡ [ a ]
81+
ℕ×ℕprf (a , b) = eq/ (Int→ℕ×ℕ (ℕ×ℕ→Int (a , b))) (a , b) (ℕ×ℕ→Int→ℕ×ℕ (a , b))
82+
83+
Int≡DiffInt : Int ≡ ℤ
84+
Int≡DiffInt = isoToPath (iso Int→ℤ ℤ→Int ℤ→Int→ℤ Int→ℤ→Int)

0 commit comments

Comments
 (0)