feat(hit/circle): prove partly that the fundamental group of the circle is int
Also add markdown files for nat and int
This commit is contained in:
parent
0a8f4f6dab
commit
9893de6194
17 changed files with 321 additions and 63 deletions
|
@ -8,9 +8,9 @@ Authors: Floris van Doorn
|
|||
Declaration of the circle
|
||||
-/
|
||||
|
||||
import .sphere types.bool types.eq
|
||||
import .sphere types.bool types.eq types.int.hott types.arrow types.equiv
|
||||
|
||||
open eq suspension bool sphere_index equiv equiv.ops
|
||||
open eq suspension bool sphere_index is_equiv equiv equiv.ops is_trunc
|
||||
|
||||
definition circle [reducible] := sphere 1
|
||||
|
||||
|
@ -130,6 +130,10 @@ namespace circle
|
|||
transport (elim_type Pbase Ploop) loop = Ploop :=
|
||||
by rewrite [tr_eq_cast_ap_fn,↑elim_type,elim_loop];apply cast_ua_fn
|
||||
|
||||
theorem elim_type_loop_inv (Pbase : Type) (Ploop : Pbase ≃ Pbase) :
|
||||
transport (elim_type Pbase Ploop) loop⁻¹ = to_inv Ploop :=
|
||||
by rewrite [tr_inv_fn,↑to_inv]; apply inv_eq_inv; apply elim_type_loop
|
||||
|
||||
definition loop_neq_idp : loop ≠ idp :=
|
||||
assume H : loop = idp,
|
||||
have H2 : Π{A : Type₁} {a : A} (p : a = a), p = idp,
|
||||
|
@ -149,4 +153,51 @@ namespace circle
|
|||
have H2 : loop = idp, from apd10 H base,
|
||||
absurd H2 loop_neq_idp
|
||||
|
||||
open int
|
||||
|
||||
protected definition code (x : circle) : Type₀ :=
|
||||
circle.elim_type_on x ℤ equiv_succ
|
||||
|
||||
definition transport_code_loop (a : ℤ) : transport code loop a = succ a :=
|
||||
ap10 !elim_type_loop a
|
||||
|
||||
definition transport_code_loop_inv (a : ℤ)
|
||||
: transport code loop⁻¹ a = pred a :=
|
||||
ap10 !elim_type_loop_inv a
|
||||
|
||||
protected definition encode {x : circle} (p : base = x) : code x :=
|
||||
transport code p (of_num 0) -- why is the explicit coercion needed here?
|
||||
|
||||
--attribute type_quotient.rec_on [unfold-c 4]
|
||||
definition circle_eq_equiv (x : circle) : (base = x) ≃ code x :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ exact encode},
|
||||
{ refine circle.rec_on x _ _,
|
||||
{ exact power loop},
|
||||
{ apply eq_of_homotopy, intro a,
|
||||
refine !arrow.arrow_transport ⬝ !transport_eq_r ⬝ _,
|
||||
rewrite [transport_code_loop_inv,power_con,succ_pred]}},
|
||||
{ refine circle.rec_on x _ _,
|
||||
{ intro a, esimp [circle.rec_on, circle.rec, base, rec2_on, rec2, base1,
|
||||
suspension.rec_on, suspension.rec, north, pushout.rec_on, pushout.rec,
|
||||
pushout.inl, type_quotient.rec_on], --simplify after #587
|
||||
apply rec_nat_on a,
|
||||
{ exact idp},
|
||||
{ intros n p,
|
||||
apply transport (λ(y : base = base), transport code y _ = _), apply power_con,
|
||||
rewrite [▸*,con_tr, transport_code_loop, ↑[encode,code] at p, p]},
|
||||
{ intros n p,
|
||||
apply transport (λ(y : base = base), transport code y _ = _),
|
||||
{ exact !power_con_inv ⬝ ap (power loop) !neg_succ⁻¹},
|
||||
rewrite [▸*,con_tr,transport_code_loop_inv, ↑[encode,code] at p, p, -neg_succ]}},
|
||||
{ apply eq_of_homotopy, intro a, apply @is_hset.elim, change is_hset ℤ, exact _}},
|
||||
--simplify after #587
|
||||
{ intro p, cases p, exact idp},
|
||||
end
|
||||
|
||||
definition base_eq_base_equiv : (base = base) ≃ ℤ :=
|
||||
circle_eq_equiv base
|
||||
|
||||
|
||||
end circle
|
||||
|
|
|
@ -212,7 +212,6 @@ namespace is_equiv
|
|||
|
||||
definition eq_inv_of_eq (p : f a = b) : a = f⁻¹ b :=
|
||||
(inv_eq_of_eq p⁻¹)⁻¹
|
||||
|
||||
end
|
||||
|
||||
--Transporting is an equivalence
|
||||
|
@ -220,8 +219,18 @@ namespace is_equiv
|
|||
is_equiv.mk _ (transport P p⁻¹) (tr_inv_tr p) (inv_tr_tr p) (tr_inv_tr_lemma p)
|
||||
|
||||
end is_equiv
|
||||
|
||||
open is_equiv
|
||||
|
||||
namespace eq
|
||||
definition tr_inv_fn {A : Type} {B : A → Type} {a a' : A} (p : a = a') :
|
||||
transport B p⁻¹ = (transport B p)⁻¹ := idp
|
||||
definition tr_inv {A : Type} {B : A → Type} {a a' : A} (p : a = a') (b : B a') :
|
||||
p⁻¹ ▸ b = (transport B p)⁻¹ b := idp
|
||||
|
||||
definition cast_inv_fn {A B : Type} (p : A = B) : cast p⁻¹ = (cast p)⁻¹ := idp
|
||||
definition cast_inv {A B : Type} (p : A = B) (b : B) : cast p⁻¹ b = (cast p)⁻¹ b := idp
|
||||
end eq
|
||||
|
||||
namespace equiv
|
||||
namespace ops
|
||||
attribute to_fun [coercion]
|
||||
|
@ -275,4 +284,9 @@ namespace equiv
|
|||
attribute equiv.refl [refl]
|
||||
attribute equiv.symm [symm]
|
||||
|
||||
namespace ops
|
||||
infixl `⬝e`:75 := equiv.trans
|
||||
postfix `⁻¹e`:(max + 1) := equiv.symm
|
||||
abbreviation erfl := @equiv.refl
|
||||
end ops
|
||||
end equiv
|
||||
|
|
|
@ -187,7 +187,7 @@ namespace eq
|
|||
definition cast [reducible] {A B : Type} (p : A = B) (a : A) : B :=
|
||||
p ▸ a
|
||||
|
||||
definition tr_inv [reducible] (P : A → Type) {x y : A} (p : x = y) (u : P y) : P x :=
|
||||
definition tr_rev [reducible] (P : A → Type) {x y : A} (p : x = y) (u : P y) : P x :=
|
||||
p⁻¹ ▸ u
|
||||
|
||||
definition ap ⦃A B : Type⦄ (f : A → B) {x y:A} (p : x = y) : f x = f y :=
|
||||
|
|
|
@ -78,7 +78,7 @@ end sum
|
|||
-- Product type
|
||||
-- ------------
|
||||
|
||||
definition pair := @prod.mk
|
||||
abbreviation pair := @prod.mk
|
||||
|
||||
namespace prod
|
||||
|
||||
|
|
|
@ -46,4 +46,12 @@ namespace arrow
|
|||
definition arrow_equiv_arrow_left (f0 : A ≃ A') : (A → B) ≃ (A' → B) :=
|
||||
arrow_equiv_arrow f0 equiv.refl
|
||||
|
||||
/- Transport -/
|
||||
|
||||
definition arrow_transport {B C : A → Type} (p : a = a') (f : B a → C a)
|
||||
: (transport (λa, B a → C a) p f) ∼ (λb, p ▸ f (p⁻¹ ▸ b)) :=
|
||||
eq.rec_on p (λx, idp)
|
||||
|
||||
|
||||
|
||||
end arrow
|
||||
|
|
|
@ -146,7 +146,7 @@ namespace eq
|
|||
equiv.mk inverse _
|
||||
|
||||
definition is_equiv_concat_left [instance] (p : a1 = a2) (a3 : A)
|
||||
: is_equiv (@concat _ a1 a2 a3 p) :=
|
||||
: is_equiv (concat p : a2 = a3 → a1 = a3) :=
|
||||
is_equiv.mk (concat p) (concat p⁻¹)
|
||||
(con_inv_cancel_left p)
|
||||
(inv_con_cancel_left p)
|
||||
|
|
|
@ -78,6 +78,10 @@ namespace is_equiv
|
|||
theorem is_hprop_is_equiv [instance] : is_hprop (is_equiv f) :=
|
||||
is_hprop_of_imp_is_contr (λ(H : is_equiv f), is_trunc_equiv_closed -2 (equiv.symm !sigma_char'))
|
||||
|
||||
definition inv_eq_inv {A B : Type} {f f' : A → B} {Hf : is_equiv f} {Hf' : is_equiv f'}
|
||||
(p : f = f') : f⁻¹ = f'⁻¹ :=
|
||||
apd011 inv p !is_hprop.elim
|
||||
|
||||
/- contractible fibers -/
|
||||
definition is_contr_fun [reducible] (f : A → B) := Π(b : B), is_contr (fiber f b)
|
||||
|
||||
|
|
|
@ -2,7 +2,7 @@
|
|||
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: int.basic
|
||||
Module: types.int.basic
|
||||
Authors: Floris van Doorn, Jeremy Avigad
|
||||
|
||||
The integers, with addition, multiplication, and subtraction. The representation of the integers is
|
||||
|
@ -146,21 +146,31 @@ int.cases_on a
|
|||
(take m, assume H : nat_abs (of_nat m) = 0, ap of_nat H)
|
||||
(take m', assume H : nat_abs (neg_succ_of_nat m') = 0, absurd H (succ_ne_zero _))
|
||||
|
||||
definition rec_nat_on {P : ℤ → Type} (z : ℤ) (H0 : P 0) (Hsucc : Π⦃n : ℕ⦄, P n → P (succ n))
|
||||
(Hpred : Π⦃n : ℕ⦄, P (- n) → P (-succ n)) : P z :=
|
||||
int.rec_on z (λn, nat.rec_on n H0 Hsucc) (λn, nat.rec_on n (Hpred H0) (λm H, Hpred H))
|
||||
|
||||
--the only computation rule of rec_nat_on which is not definitional
|
||||
definition rec_nat_on_neg {P : ℤ → Type} (n : nat) (H0 : P zero)
|
||||
(Hsucc : Π⦃n : nat⦄, P n → P (succ n)) (Hpred : Π⦃n : nat⦄, P (- n) → P (-succ n))
|
||||
: rec_nat_on (-succ n) H0 Hsucc Hpred = Hpred (rec_nat_on (-n) H0 Hsucc Hpred) :=
|
||||
nat.rec_on n rfl (λn H, rfl)
|
||||
|
||||
/- int is a quotient of ordered pairs of natural numbers -/
|
||||
|
||||
protected definition equiv (p q : ℕ × ℕ) : Type₀ := pr1 p + pr2 q = pr2 p + pr1 q
|
||||
definition int_equiv (p q : ℕ × ℕ) : Type₀ := pr1 p + pr2 q = pr2 p + pr1 q
|
||||
|
||||
local infix `≡` := int.equiv
|
||||
local infix `≡` := int_equiv
|
||||
|
||||
protected theorem equiv.refl [refl] {p : ℕ × ℕ} : p ≡ p := !add.comm
|
||||
protected theorem int_equiv.refl [refl] {p : ℕ × ℕ} : p ≡ p := !add.comm
|
||||
|
||||
protected theorem equiv.symm [symm] {p q : ℕ × ℕ} (H : p ≡ q) : q ≡ p :=
|
||||
protected theorem int_equiv.symm [symm] {p q : ℕ × ℕ} (H : p ≡ q) : q ≡ p :=
|
||||
calc
|
||||
pr1 q + pr2 p = pr2 p + pr1 q : !add.comm
|
||||
... = pr1 p + pr2 q : H⁻¹
|
||||
... = pr2 q + pr1 p : !add.comm
|
||||
|
||||
protected theorem equiv.trans [trans] {p q r : ℕ × ℕ} (H1 : p ≡ q) (H2 : q ≡ r) : p ≡ r :=
|
||||
protected theorem int_equiv.trans [trans] {p q r : ℕ × ℕ} (H1 : p ≡ q) (H2 : q ≡ r) : p ≡ r :=
|
||||
have H3 : pr1 p + pr2 r + pr2 q = pr2 p + pr1 r + pr2 q, from
|
||||
calc
|
||||
pr1 p + pr2 r + pr2 q = pr1 p + pr2 q + pr2 r : by exact sorry
|
||||
|
@ -170,10 +180,10 @@ have H3 : pr1 p + pr2 r + pr2 q = pr2 p + pr1 r + pr2 q, from
|
|||
... = pr2 p + pr1 r + pr2 q : by exact sorry,
|
||||
show pr1 p + pr2 r = pr2 p + pr1 r, from add.cancel_right H3
|
||||
|
||||
protected definition equiv_equiv : is_equivalence equiv :=
|
||||
is_equivalence.mk @equiv.refl @equiv.symm @equiv.trans
|
||||
protected definition int_equiv_int_equiv : is_equivalence int_equiv :=
|
||||
is_equivalence.mk @int_equiv.refl @int_equiv.symm @int_equiv.trans
|
||||
|
||||
protected definition equiv_cases {p q : ℕ × ℕ} (H : equiv p q) :
|
||||
protected definition int_equiv_cases {p q : ℕ × ℕ} (H : int_equiv p q) :
|
||||
(pr1 p ≥ pr2 p × pr1 q ≥ pr2 q) ⊎ (pr1 p < pr2 p × pr1 q < pr2 q) :=
|
||||
sum.rec_on (@le_or_gt (pr2 p) (pr1 p))
|
||||
(assume H1: pr1 p ≥ pr2 p,
|
||||
|
@ -183,7 +193,7 @@ sum.rec_on (@le_or_gt (pr2 p) (pr1 p))
|
|||
have H2 : pr2 p + pr1 q < pr2 p + pr2 q, from H ▸ add_lt_add_right H1 (pr2 q),
|
||||
sum.inr (pair H1 (lt_of_add_lt_add_left H2)))
|
||||
|
||||
protected definition equiv_of_eq {p q : ℕ × ℕ} (H : p = q) : p ≡ q := H ▸ equiv.refl
|
||||
protected definition int_equiv_of_eq {p q : ℕ × ℕ} (H : p = q) : p ≡ q := H ▸ int_equiv.refl
|
||||
|
||||
/- the representation and abstraction functions -/
|
||||
|
||||
|
@ -220,15 +230,15 @@ sum.rec_on (@le_or_gt n m)
|
|||
definition repr_abstr (p : ℕ × ℕ) : repr (abstr p) ≡ p :=
|
||||
!prod.eta ▸ !repr_sub_nat_nat
|
||||
|
||||
definition abstr_eq {p q : ℕ × ℕ} (Hequiv : p ≡ q) : abstr p = abstr q :=
|
||||
sum.rec_on (equiv_cases Hequiv)
|
||||
definition abstr_eq {p q : ℕ × ℕ} (Hint_equiv : p ≡ q) : abstr p = abstr q :=
|
||||
sum.rec_on (int_equiv_cases Hint_equiv)
|
||||
(assume H2,
|
||||
have H3 : pr1 p ≥ pr2 p, from prod.pr1 H2,
|
||||
have H4 : pr1 q ≥ pr2 q, from prod.pr2 H2,
|
||||
have H5 : pr1 p = pr1 q - pr2 q + pr2 p, from
|
||||
calc
|
||||
pr1 p = pr1 p + pr2 q - pr2 q : add_sub_cancel
|
||||
... = pr2 p + pr1 q - pr2 q : by rewrite [↑int.equiv at Hequiv,Hequiv]
|
||||
... = pr2 p + pr1 q - pr2 q : by rewrite [↑int_equiv at Hint_equiv,Hint_equiv]
|
||||
... = pr2 p + (pr1 q - pr2 q) : add_sub_assoc H4
|
||||
... = pr1 q - pr2 q + pr2 p : add.comm,
|
||||
have H6 : pr1 p - pr2 p = pr1 q - pr2 q, from
|
||||
|
@ -242,7 +252,7 @@ sum.rec_on (equiv_cases Hequiv)
|
|||
have H5 : pr2 p = pr2 q - pr1 q + pr1 p, from
|
||||
calc
|
||||
pr2 p = pr2 p + pr1 q - pr1 q : add_sub_cancel
|
||||
... = pr1 p + pr2 q - pr1 q : by rewrite [↑int.equiv at Hequiv,Hequiv]
|
||||
... = pr1 p + pr2 q - pr1 q : by rewrite [↑int_equiv at Hint_equiv,Hint_equiv]
|
||||
... = pr1 p + (pr2 q - pr1 q) : add_sub_assoc (le_of_lt H4)
|
||||
... = pr2 q - pr1 q + pr1 p : add.comm,
|
||||
have H6 : pr2 p - pr1 p = pr2 q - pr1 q, from
|
||||
|
@ -251,20 +261,20 @@ sum.rec_on (equiv_cases Hequiv)
|
|||
... = pr2 q - pr1 q : add_sub_cancel,
|
||||
abstr_of_lt H3 ⬝ ap neg_succ_of_nat (ap pred H6)⬝ (abstr_of_lt H4)⁻¹)
|
||||
|
||||
definition equiv_iff (p q : ℕ × ℕ) : (p ≡ q) ↔ ((p ≡ p) × (q ≡ q) × (abstr p = abstr q)) :=
|
||||
definition int_equiv_iff (p q : ℕ × ℕ) : (p ≡ q) ↔ ((p ≡ p) × (q ≡ q) × (abstr p = abstr q)) :=
|
||||
iff.intro
|
||||
(assume H : equiv p q,
|
||||
pair !equiv.refl (pair !equiv.refl (abstr_eq H)))
|
||||
(assume H : equiv p p × equiv q q × abstr p = abstr q,
|
||||
(assume H : int_equiv p q,
|
||||
pair !int_equiv.refl (pair !int_equiv.refl (abstr_eq H)))
|
||||
(assume H : int_equiv p p × int_equiv q q × abstr p = abstr q,
|
||||
have H1 : abstr p = abstr q, from prod.pr2 (prod.pr2 H),
|
||||
equiv.trans (H1 ▸ equiv.symm (repr_abstr p)) (repr_abstr q))
|
||||
int_equiv.trans (H1 ▸ int_equiv.symm (repr_abstr p)) (repr_abstr q))
|
||||
|
||||
definition eq_abstr_of_equiv_repr {a : ℤ} {p : ℕ × ℕ} (Hequiv : repr a ≡ p) : a = abstr p :=
|
||||
definition eq_abstr_of_int_equiv_repr {a : ℤ} {p : ℕ × ℕ} (Hint_equiv : repr a ≡ p) : a = abstr p :=
|
||||
calc
|
||||
a = abstr (repr a) : abstr_repr
|
||||
... = abstr p : abstr_eq Hequiv
|
||||
... = abstr p : abstr_eq Hint_equiv
|
||||
|
||||
definition eq_of_repr_equiv_repr {a b : ℤ} (H : repr a ≡ repr b) : a = b :=
|
||||
definition eq_of_repr_int_equiv_repr {a b : ℤ} (H : repr a ≡ repr b) : a = b :=
|
||||
calc
|
||||
a = abstr (repr a) : abstr_repr
|
||||
... = abstr (repr b) : abstr_eq H
|
||||
|
@ -322,9 +332,9 @@ definition repr_add (a b : ℤ) : repr (add a b) ≡ padd (repr a) (repr b) :=
|
|||
int.cases_on a
|
||||
(take m,
|
||||
int.cases_on b
|
||||
(take n, !equiv.refl)
|
||||
(take n, !int_equiv.refl)
|
||||
(take n',
|
||||
have H1 : equiv (repr (add (of_nat m) (neg_succ_of_nat n'))) (m, succ n'),
|
||||
have H1 : int_equiv (repr (add (of_nat m) (neg_succ_of_nat n'))) (m, succ n'),
|
||||
from !repr_sub_nat_nat,
|
||||
have H2 : padd (repr (of_nat m)) (repr (neg_succ_of_nat n')) = (m, 0 + succ n'),
|
||||
from rfl,
|
||||
|
@ -332,7 +342,7 @@ int.cases_on a
|
|||
(take m',
|
||||
int.cases_on b
|
||||
(take n,
|
||||
have H1 : equiv (repr (add (neg_succ_of_nat m') (of_nat n))) (n, succ m'),
|
||||
have H1 : int_equiv (repr (add (neg_succ_of_nat m') (of_nat n))) (n, succ m'),
|
||||
from !repr_sub_nat_nat,
|
||||
have H2 : padd (repr (neg_succ_of_nat m')) (repr (of_nat n)) = (0 + n, succ m'),
|
||||
from rfl,
|
||||
|
@ -362,25 +372,25 @@ calc
|
|||
|
||||
definition add.comm (a b : ℤ) : a + b = b + a :=
|
||||
begin
|
||||
apply eq_of_repr_equiv_repr,
|
||||
apply equiv.trans,
|
||||
apply eq_of_repr_int_equiv_repr,
|
||||
apply int_equiv.trans,
|
||||
apply repr_add,
|
||||
apply equiv.symm,
|
||||
apply int_equiv.symm,
|
||||
apply eq.subst (padd_comm (repr b) (repr a)),
|
||||
apply repr_add
|
||||
end
|
||||
|
||||
definition add.assoc (a b c : ℤ) : a + b + c = a + (b + c) :=
|
||||
assert H1 : repr (a + b + c) ≡ padd (padd (repr a) (repr b)) (repr c), from
|
||||
equiv.trans (repr_add (a + b) c) (padd_congr !repr_add !equiv.refl),
|
||||
int_equiv.trans (repr_add (a + b) c) (padd_congr !repr_add !int_equiv.refl),
|
||||
assert H2 : repr (a + (b + c)) ≡ padd (repr a) (padd (repr b) (repr c)), from
|
||||
equiv.trans (repr_add a (b + c)) (padd_congr !equiv.refl !repr_add),
|
||||
int_equiv.trans (repr_add a (b + c)) (padd_congr !int_equiv.refl !repr_add),
|
||||
begin
|
||||
apply eq_of_repr_equiv_repr,
|
||||
apply equiv.trans,
|
||||
apply eq_of_repr_int_equiv_repr,
|
||||
apply int_equiv.trans,
|
||||
apply H1,
|
||||
apply eq.subst (padd_assoc _ _ _)⁻¹,
|
||||
apply equiv.symm,
|
||||
apply int_equiv.symm,
|
||||
apply H2
|
||||
end
|
||||
|
||||
|
@ -424,7 +434,7 @@ have H : repr (-a + a) ≡ repr 0, from
|
|||
repr (-a + a) ≡ padd (repr (neg a)) (repr a) : repr_add
|
||||
... ≡ padd (pneg (repr a)) (repr a) : sorry
|
||||
... ≡ repr 0 : padd_pneg,
|
||||
eq_of_repr_equiv_repr H
|
||||
eq_of_repr_int_equiv_repr H
|
||||
|
||||
/- nat abs -/
|
||||
|
||||
|
@ -506,7 +516,7 @@ int.cases_on a
|
|||
... = (succ m' * succ n', 0) : zero_mul
|
||||
... = repr (mul (neg_succ_of_nat m') (neg_succ_of_nat n')) : rfl)⁻¹))
|
||||
|
||||
definition equiv_mul_prep {xa ya xb yb xn yn xm ym : ℕ}
|
||||
definition int_equiv_mul_prep {xa ya xb yb xn yn xm ym : ℕ}
|
||||
(H1 : xa + yb = ya + xb) (H2 : xn + ym = yn + xm)
|
||||
: xa * xn + ya * yn + (xb * ym + yb * xm) = xa * yn + ya * xn + (xb * xm + yb * ym) :=
|
||||
have H3 : xa * xn + ya * yn + (xb * ym + yb * xm) + (yb * xn + xb * yn + (xb * xn + yb * yn))
|
||||
|
@ -524,7 +534,7 @@ have H3 : xa * xn + ya * yn + (xb * ym + yb * xm) + (yb * xn + xb * yn + (xb * x
|
|||
nat.add.cancel_right H3
|
||||
|
||||
definition pmul_congr {p p' q q' : ℕ × ℕ} (H1 : p ≡ p') (H2 : q ≡ q') : pmul p q ≡ pmul p' q' :=
|
||||
equiv_mul_prep H1 H2
|
||||
int_equiv_mul_prep H1 H2
|
||||
|
||||
definition pmul_comm (p q : ℕ × ℕ) : pmul p q = pmul q p :=
|
||||
calc
|
||||
|
@ -536,26 +546,26 @@ calc
|
|||
... = (pr1 q * pr1 p + pr2 q * pr2 p, pr1 q * pr2 p + pr2 q * pr1 p) : nat.add.comm
|
||||
|
||||
definition mul.comm (a b : ℤ) : a * b = b * a :=
|
||||
eq_of_repr_equiv_repr
|
||||
eq_of_repr_int_equiv_repr
|
||||
((calc
|
||||
repr (a * b) = pmul (repr a) (repr b) : repr_mul
|
||||
... = pmul (repr b) (repr a) : pmul_comm
|
||||
... = repr (b * a) : repr_mul) ▸ !equiv.refl)
|
||||
... = repr (b * a) : repr_mul) ▸ !int_equiv.refl)
|
||||
|
||||
definition pmul_assoc (p q r: ℕ × ℕ) : pmul (pmul p q) r = pmul p (pmul q r) :=
|
||||
by exact sorry
|
||||
|
||||
definition mul.assoc (a b c : ℤ) : (a * b) * c = a * (b * c) :=
|
||||
eq_of_repr_equiv_repr
|
||||
eq_of_repr_int_equiv_repr
|
||||
((calc
|
||||
repr (a * b * c) = pmul (repr (a * b)) (repr c) : repr_mul
|
||||
... = pmul (pmul (repr a) (repr b)) (repr c) : repr_mul
|
||||
... = pmul (repr a) (pmul (repr b) (repr c)) : pmul_assoc
|
||||
... = pmul (repr a) (repr (b * c)) : repr_mul
|
||||
... = repr (a * (b * c)) : repr_mul) ▸ !equiv.refl)
|
||||
... = repr (a * (b * c)) : repr_mul) ▸ !int_equiv.refl)
|
||||
|
||||
definition mul_one (a : ℤ) : a * 1 = a :=
|
||||
eq_of_repr_equiv_repr (equiv_of_eq
|
||||
eq_of_repr_int_equiv_repr (int_equiv_of_eq
|
||||
((calc
|
||||
repr (a * 1) = pmul (repr a) (repr 1) : repr_mul
|
||||
... = (pr1 (repr a), pr2 (repr a)) : by exact sorry
|
||||
|
@ -565,14 +575,14 @@ definition one_mul (a : ℤ) : 1 * a = a :=
|
|||
mul.comm a 1 ▸ mul_one a
|
||||
|
||||
definition mul.right_distrib (a b c : ℤ) : (a + b) * c = a * c + b * c :=
|
||||
eq_of_repr_equiv_repr
|
||||
eq_of_repr_int_equiv_repr
|
||||
(calc
|
||||
repr ((a + b) * c) = pmul (repr (a + b)) (repr c) : repr_mul
|
||||
... ≡ pmul (padd (repr a) (repr b)) (repr c) : pmul_congr !repr_add equiv.refl
|
||||
... ≡ pmul (padd (repr a) (repr b)) (repr c) : pmul_congr !repr_add int_equiv.refl
|
||||
... = padd (pmul (repr a) (repr c)) (pmul (repr b) (repr c)) : by exact sorry
|
||||
... = padd (repr (a * c)) (pmul (repr b) (repr c)) : {(repr_mul a c)⁻¹}
|
||||
... = padd (repr (a * c)) (repr (b * c)) : repr_mul
|
||||
... ≡ repr (a * c + b * c) : equiv.symm !repr_add)
|
||||
... ≡ repr (a * c + b * c) : int_equiv.symm !repr_add)
|
||||
|
||||
definition mul.left_distrib (a b c : ℤ) : a * (b + c) = a * b + a * c :=
|
||||
calc
|
||||
|
|
9
hott/types/int/default.hlean
Normal file
9
hott/types/int/default.hlean
Normal file
|
@ -0,0 +1,9 @@
|
|||
/-
|
||||
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: types.int.default
|
||||
Authors: Floris van Doorn
|
||||
-/
|
||||
|
||||
import .basic
|
130
hott/types/int/hott.hlean
Normal file
130
hott/types/int/hott.hlean
Normal file
|
@ -0,0 +1,130 @@
|
|||
/-
|
||||
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: types.int.hott
|
||||
Author: Floris van Doorn
|
||||
|
||||
Theorems about the integers specific to HoTT
|
||||
-/
|
||||
|
||||
import .basic types.eq
|
||||
open core is_equiv equiv nat equiv.ops
|
||||
|
||||
namespace int
|
||||
|
||||
definition succ (a : ℤ) := a + (succ zero)
|
||||
definition pred (a : ℤ) := a - (succ zero)
|
||||
definition pred_succ (a : ℤ) : pred (succ a) = a := !sub_add_cancel
|
||||
definition succ_pred (a : ℤ) : succ (pred a) = a := !add_sub_cancel
|
||||
definition neg_succ (a : ℤ) : -succ a = pred (-a) :=
|
||||
by rewrite [↑succ,neg_add]
|
||||
definition succ_neg_succ (a : ℤ) : succ (-succ a) = -a :=
|
||||
by rewrite [neg_succ,succ_pred]
|
||||
|
||||
definition is_equiv_succ [instance] : is_equiv succ :=
|
||||
adjointify succ pred (λa, !add_sub_cancel) (λa, !sub_add_cancel)
|
||||
definition equiv_succ : ℤ ≃ ℤ := equiv.mk succ _
|
||||
|
||||
definition is_equiv_neg [instance] : is_equiv neg :=
|
||||
adjointify neg neg (λx, !neg_neg) (λa, !neg_neg)
|
||||
definition equiv_neg : ℤ ≃ ℤ := equiv.mk neg _
|
||||
|
||||
definition iterate {A : Type} (f : A ≃ A) (a : ℤ) : A ≃ A :=
|
||||
rec_nat_on a erfl
|
||||
(λb g, f ⬝e g)
|
||||
(λb g, g ⬝e f⁻¹e)
|
||||
|
||||
-- definition iterate_trans {A : Type} (f : A ≃ A) (a : ℤ)
|
||||
-- : iterate f a ⬝e f = iterate f (a + 1) :=
|
||||
-- sorry
|
||||
|
||||
-- definition trans_iterate {A : Type} (f : A ≃ A) (a : ℤ)
|
||||
-- : f ⬝e iterate f a = iterate f (a + 1) :=
|
||||
-- sorry
|
||||
|
||||
-- definition iterate_trans_symm {A : Type} (f : A ≃ A) (a : ℤ)
|
||||
-- : iterate f a ⬝e f⁻¹e = iterate f (a - 1) :=
|
||||
-- sorry
|
||||
|
||||
-- definition symm_trans_iterate {A : Type} (f : A ≃ A) (a : ℤ)
|
||||
-- : f⁻¹e ⬝e iterate f a = iterate f (a - 1) :=
|
||||
-- sorry
|
||||
|
||||
-- definition iterate_neg {A : Type} (f : A ≃ A) (a : ℤ)
|
||||
-- : iterate f (-a) = (iterate f a)⁻¹e :=
|
||||
-- rec_nat_on a idp
|
||||
-- (λn p, calc
|
||||
-- iterate f (-succ n) = iterate f (-n) ⬝e f⁻¹e : rec_nat_on_neg
|
||||
-- ... = (iterate f n)⁻¹e ⬝e f⁻¹e : by rewrite p
|
||||
-- ... = (f ⬝e iterate f n)⁻¹e : sorry
|
||||
-- ... = (iterate f (succ n))⁻¹e : idp)
|
||||
-- sorry
|
||||
|
||||
-- definition iterate_add {A : Type} (f : A ≃ A) (a b : ℤ)
|
||||
-- : iterate f (a + b) = equiv.trans (iterate f a) (iterate f b) :=
|
||||
-- sorry
|
||||
|
||||
-- definition iterate_sub {A : Type} (f : A ≃ A) (a b : ℤ)
|
||||
-- : iterate f (a - b) = equiv.trans (iterate f a) (equiv.symm (iterate f b)) :=
|
||||
-- sorry
|
||||
|
||||
-- definition iterate_mul {A : Type} (f : A ≃ A) (a b : ℤ)
|
||||
-- : iterate f (a * b) = iterate (iterate f a) b :=
|
||||
-- sorry
|
||||
|
||||
end int open int
|
||||
|
||||
|
||||
|
||||
namespace eq
|
||||
variables {A : Type} {a : A} (p : a = a) (b : ℤ)
|
||||
definition power : a = a :=
|
||||
rec_nat_on b idp
|
||||
(λc q, q ⬝ p)
|
||||
(λc q, q ⬝ p⁻¹)
|
||||
definition power_neg_succ (n : ℕ) : power p (-succ n) = power p (-n) ⬝ p⁻¹ :=
|
||||
begin
|
||||
rewrite [↑power], exact sorry
|
||||
end
|
||||
-- rec_nat_on_neg idp
|
||||
-- (λc q, q ⬝ p)
|
||||
-- (λc q, q ⬝ p⁻¹) n
|
||||
|
||||
--iterate (equiv_eq_closed_right p a) b idp
|
||||
set_option pp.coercions true
|
||||
-- attribute nat.add int.add int.of_num nat.of_num int.succ [constructor]
|
||||
attribute rec_nat_on [unfold-c 2]
|
||||
|
||||
definition power_con : power p b ⬝ p = power p (succ b) :=
|
||||
rec_nat_on b
|
||||
idp
|
||||
(λn IH, idp)
|
||||
(λn IH, calc
|
||||
power p (-succ n) ⬝ p = (power p (-n) ⬝ p⁻¹) ⬝ p : {!rec_nat_on_neg}
|
||||
... = power p (-n) : inv_con_cancel_right
|
||||
... = power p (succ (-succ n)) : {(succ_neg_succ n)⁻¹})
|
||||
|
||||
definition con_power : p ⬝ power p b = power p (succ b) :=
|
||||
rec_nat_on b
|
||||
(by rewrite ↑[power];exact !idp_con⁻¹)
|
||||
(λn IH, calc
|
||||
p ⬝ power p (succ n) = (p ⬝ power p n) ⬝ p : con.assoc
|
||||
... = power p (succ (succ n)) : by rewrite IH)
|
||||
(λn IH, calc
|
||||
p ⬝ power p (-succ n) = p ⬝ (power p (-n) ⬝ p⁻¹) : sorry
|
||||
... = (p ⬝ power p (-n)) ⬝ p⁻¹ : con.assoc
|
||||
... = power p (succ (-n)) ⬝ p⁻¹ : by rewrite IH
|
||||
... = power p (succ (-succ n)) : sorry)
|
||||
|
||||
definition power_con_inv : power p b ⬝ p⁻¹ = power p (pred b) :=
|
||||
rec_nat_on b sorry
|
||||
sorry
|
||||
sorry
|
||||
|
||||
-- definition inv_con_power : p⁻¹ ⬝ power p b = power p (pred b) :=
|
||||
-- rec_nat_on b sorry
|
||||
-- sorry
|
||||
-- sorry
|
||||
|
||||
end eq
|
7
hott/types/int/int.md
Normal file
7
hott/types/int/int.md
Normal file
|
@ -0,0 +1,7 @@
|
|||
types.int
|
||||
=========
|
||||
|
||||
The integers. Note: most of these files are ported from the standard library. If anything needs to be changed, it is probably a good idea to change it in the standard library and then port the file again (see also [script/port.pl](../../script/port.pl)).
|
||||
|
||||
* [basic](basic.hlean) : the integers, with basic operations
|
||||
* [hott](hott.hlean) : facts about the integers specific to the HoTT library
|
8
hott/types/nat/nat.md
Normal file
8
hott/types/nat/nat.md
Normal file
|
@ -0,0 +1,8 @@
|
|||
types.nat
|
||||
=========
|
||||
|
||||
The natural numbers. Note: all these files are ported from the standard library. If anything needs to be changed, it is probably a good idea to change it in the standard library and then port the file again (see also [script/port.pl](../../script/port.pl)).
|
||||
|
||||
* [basic](basic.hlean) : the natural numbers, with succ, pred, addition, and multiplication
|
||||
* [order](order.hlean) : less-than, less-then-or-equal, etc.
|
||||
* [sub](sub.hlean) : subtraction, and distance
|
|
@ -126,16 +126,16 @@ namespace pi
|
|||
unfold pi_functor, unfold function.compose, unfold function.id,
|
||||
begin
|
||||
intro a',
|
||||
apply (tr_inv _ (adj f0 a')),
|
||||
apply (tr_rev _ (adj f0 a')),
|
||||
apply (transport (λx, f1 a' x = h a') (transport_compose B f0 (left_inv f0 a') _)),
|
||||
apply (tr_inv (λx, x = h a') (fn_tr_eq_tr_fn _ f1 _)), unfold function.compose,
|
||||
apply (tr_inv (λx, left_inv f0 a' ▸ x = h a') (right_inv (f1 _) _)), unfold function.id,
|
||||
apply (tr_rev (λx, x = h a') (fn_tr_eq_tr_fn _ f1 _)), unfold function.compose,
|
||||
apply (tr_rev (λx, left_inv f0 a' ▸ x = h a') (right_inv (f1 _) _)), unfold function.id,
|
||||
apply apd
|
||||
end,
|
||||
begin
|
||||
intro h,
|
||||
apply eq_of_homotopy, intro a,
|
||||
apply (tr_inv (λx, right_inv f0 a ▸ x = h a) (left_inv (f1 _) _)), unfold function.id,
|
||||
apply (tr_rev (λx, right_inv f0 a ▸ x = h a) (left_inv (f1 _) _)), unfold function.id,
|
||||
apply apd
|
||||
end
|
||||
end
|
||||
|
|
|
@ -145,21 +145,24 @@ end is_trunc open is_trunc
|
|||
namespace trunc
|
||||
variable {A : Type}
|
||||
|
||||
definition trunc_eq_type (n : trunc_index) (aa aa' : trunc n.+1 A) : n-Type :=
|
||||
protected definition code (n : trunc_index) (aa aa' : trunc n.+1 A) : n-Type :=
|
||||
trunc.rec_on aa (λa, trunc.rec_on aa' (λa', trunctype.mk' n (trunc n (a = a'))))
|
||||
|
||||
-- protected definition encode (n : trunc_index) (aa aa' : trunc n.+1 A) : aa = aa' → code n aa aa' :=
|
||||
-- trunc.rec_on aa (λa, trunc.rec_on aa' (λa' p, _))
|
||||
|
||||
definition trunc_eq_equiv (n : trunc_index) (aa aa' : trunc n.+1 A)
|
||||
: aa = aa' ≃ trunc_eq_type n aa aa' :=
|
||||
: aa = aa' ≃ code n aa aa' :=
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro p, cases p, apply (trunc.rec_on aa),
|
||||
intro a, esimp [trunc_eq_type,trunc.rec_on], exact (tr idp)},
|
||||
intro a, esimp [code,trunc.rec_on], exact (tr idp)},
|
||||
{ eapply (trunc.rec_on aa'), eapply (trunc.rec_on aa),
|
||||
intro a a' x, esimp [trunc_eq_type, trunc.rec_on] at x,
|
||||
intro a a' x, esimp [code, trunc.rec_on] at x,
|
||||
apply (trunc.rec_on x), intro p, exact (ap tr p)},
|
||||
{
|
||||
-- apply (trunc.rec_on aa'), apply (trunc.rec_on aa),
|
||||
-- intro a a' x, esimp [trunc_eq_type, trunc.rec_on] at x,
|
||||
-- intro a a' x, esimp [code, trunc.rec_on] at x,
|
||||
-- apply (trunc.rec_on x), intro p,
|
||||
-- cases p, esimp [trunc.rec_on,eq.cases_on,compose,id], -- apply idp --?
|
||||
apply sorry},
|
||||
|
|
|
@ -16,3 +16,8 @@ Various datatypes.
|
|||
* [function](function.hlean): embeddings, (split) surjections, retractions
|
||||
* [trunc](trunc.hlean): truncation levels, n-Types, truncation
|
||||
* [W](W.hlean): W-types (not loaded by default)
|
||||
|
||||
Subfolders:
|
||||
|
||||
* [nat](nat/nat.md)
|
||||
* [int](int/int.md)
|
|
@ -2,7 +2,7 @@
|
|||
Copyright (c) 2014 Floris van Doorn. All rights reserved.
|
||||
Released under Apache 2.0 license as described in the file LICENSE.
|
||||
|
||||
Module: int.basic
|
||||
Module: data.int.basic
|
||||
Authors: Floris van Doorn, Jeremy Avigad
|
||||
|
||||
The integers, with addition, multiplication, and subtraction. The representation of the integers is
|
||||
|
@ -116,7 +116,7 @@ int.cases_on a
|
|||
(if H : m' = n' then inl (congr_arg neg_succ_of_nat H) else
|
||||
inr (take H1, H (neg_succ_of_nat.inj H1)))))
|
||||
|
||||
theorem of_nat_add_of_nat (n m : nat) : of_nat n + of_nat m = #nat n + m := rfl
|
||||
theorem of_nat_add_of_nat (n m : nat) : #int n + m = #nat n + m := rfl
|
||||
|
||||
theorem of_nat_succ (n : ℕ) : of_nat (succ n) = of_nat n + 1 := rfl
|
||||
|
||||
|
@ -148,6 +148,16 @@ int.cases_on a
|
|||
(take m, assume H : nat_abs (of_nat m) = 0, congr_arg of_nat H)
|
||||
(take m', assume H : nat_abs (neg_succ_of_nat m') = 0, absurd H (succ_ne_zero _))
|
||||
|
||||
definition rec_nat_on {P : ℤ → Type} (z : ℤ) (H0 : P 0) (Hsucc : Π⦃n : ℕ⦄, P n → P (succ n))
|
||||
(Hpred : Π⦃n : ℕ⦄, P (- n) → P (-succ n)) : P z :=
|
||||
int.rec_on z (λn, nat.rec_on n H0 Hsucc) (λn, nat.rec_on n (Hpred H0) (λm H, Hpred H))
|
||||
|
||||
--the only computation rule of rec_nat_on which is not definitional
|
||||
definition rec_nat_on_neg {P : ℤ → Type} (n : nat) (H0 : P 0)
|
||||
(Hsucc : Π⦃n : nat⦄, P n → P (succ n)) (Hpred : Π⦃n : nat⦄, P (- n) → P (-succ n))
|
||||
: rec_nat_on (-succ n) H0 Hsucc Hpred = Hpred (rec_nat_on (-n) H0 Hsucc Hpred) :=
|
||||
nat.rec_on n rfl (λn H, rfl)
|
||||
|
||||
/- int is a quotient of ordered pairs of natural numbers -/
|
||||
|
||||
protected definition equiv (p q : ℕ × ℕ) : Prop := pr1 p + pr2 q = pr2 p + pr1 q
|
||||
|
|
|
@ -1,5 +1,4 @@
|
|||
Prop:Type
|
||||
theorem:definition
|
||||
by simp;by exact sorry
|
||||
|
||||
true:unit
|
||||
|
|
Loading…
Reference in a new issue