Added old proofs in new style (missing: Z' is hset)

This commit is contained in:
Niels 2017-08-04 20:42:05 +02:00
parent 8cf115c9ab
commit efec2e88f8
1 changed files with 258 additions and 441 deletions

View File

@ -76,7 +76,7 @@ Module Export Ints.
End Z_recursion. End Z_recursion.
Instance FSet_recursion : HitRecursion Z := Instance Z_recursion : HitRecursion Z :=
{ {
indTy := _; recTy := _; indTy := _; recTy := _;
H_inductor := Z_ind; H_recursor := Z_rec }. H_inductor := Z_ind; H_recursor := Z_rec }.
@ -477,9 +477,6 @@ End initial_Z.
(*
Module Export AltInts. Module Export AltInts.
Private Inductive Z' : Type0 := Private Inductive Z' : Type0 :=
@ -488,443 +485,263 @@ Module Export AltInts.
Axiom path : positive 0 = negative 0. Axiom path : positive 0 = negative 0.
Fixpoint Z'_ind Section AltIntsInd.
(P : Z' -> Type) Variable (P : Z' -> Type)
(po : forall (x : nat), P (positive x)) (po : forall (x : nat), P (positive x))
(ne : forall (x : nat), P (negative x)) (ne : forall (x : nat), P (negative x))
(i : path # (po 0) = ne 0) (i : path # (po 0) = ne 0).
(x : Z')
{struct x} Fixpoint Z'_ind (x : Z') {struct x} : P x
: P x :=
:= (match x return _ -> P x with
(match x return _ -> P x with | positive n => fun _ => po n
| positive n => fun _ => po n | negative n => fun _ => ne n
| negative n => fun _ => ne n end) i.
end) i.
Axiom Z'_ind_beta_path : apD Z'_ind path = i.
End AltIntsInd.
Section AltIntsRec.
Variable (P : Type)
(po : nat -> P)
(pn : nat -> P)
(i : po 0 = pn 0).
Definition Z'_rec : Z' -> P.
Proof.
simple refine (Z'_ind _ _ _ _) ; simpl.
- apply po.
- apply pn.
- refine (transport_const _ _ @ i).
Defined.
Definition Z'_rec_beta_path : ap Z'_rec path = i.
Proof.
unfold Z_rec.
eapply (cancelL (transport_const path _)).
simple refine ((apD_const _ _)^ @ _).
apply Z'_ind_beta_path.
Defined.
End AltIntsRec.
Instance Z'_recursion : HitRecursion Z' :=
{
indTy := _; recTy := _;
H_inductor := Z'_ind; H_recursor := Z'_rec
}.
Axiom Z'_ind_beta_inv1 : forall
(P : Z' -> Type)
(po : forall (x : nat), P (positive x))
(ne : forall (x : nat), P (negative x))
(i : path # (po 0) = ne 0)
, apD (Z'_ind P po ne i) path = i.
End AltInts. End AltInts.
Definition succ_Z' : Z' -> Z'. Section Isomorphic.
Proof.
refine (Z'_ind _ _ _ _). Definition succ_Z' : Z' -> Z'.
Unshelve. Proof.
Focus 2. simple refine (Z'_rec _ _ _ _).
intro n. - apply (fun n => positive (S n)).
apply (positive (S n)). - induction 1 as [ | n].
* apply (positive 1).
Focus 2. * apply (negative n).
intro n. - reflexivity.
induction n. Defined.
apply (positive 1).
Definition pred_Z' : Z' -> Z'.
apply (negative n). Proof.
simple refine (Z'_rec _ _ _ _).
simpl. - induction 1 as [ | n].
rewrite transport_const. * apply (negative 1).
reflexivity. * apply (positive n).
Defined. - intro n.
apply (negative (S n)).
Definition pred_Z' : Z' -> Z'. - reflexivity.
Proof. Defined.
refine (Z'_ind _ _ _ _).
Unshelve. Fixpoint Nat_to_Pos (n : nat) : Int.Pos :=
Focus 2. match n with
intro n. | 0 => Int.one
induction n. | S k => succ_pos (Nat_to_Pos k)
apply (negative 1). end.
apply (positive n). Definition Z'_to_Int : Z' -> Int.
Proof.
Focus 2. simple refine (Z'_rec _ _ _ _).
intro n. - induction 1 as [ | n IHn].
apply (negative (S n)). apply (Int.zero).
apply (succ_int IHn).
simpl. - induction 1 as [ | n IHn].
rewrite transport_const. apply (Int.zero).
reflexivity. apply (pred_int IHn).
Defined. - reflexivity.
Defined.
Fixpoint Nat_to_Pos (n : nat) : Pos := Definition Pos_to_Nat : Int.Pos -> nat.
match n with Proof.
| 0 => Int.one induction 1 as [ | n IHn].
| S k => succ_pos (Nat_to_Pos k) - apply 1.
end. - apply (S IHn).
Defined.
Definition Z'_to_Int : Z' -> Int.
Proof. Definition Int_to_Z' (x : Int) : Z'.
refine (Z'_ind _ _ _ _). Proof.
Unshelve. induction x as [p | | p].
apply (negative (Pos_to_Nat p)).
Focus 2. apply (positive 0).
intro x. apply (positive (Pos_to_Nat p)).
induction x. Defined.
apply (Int.zero).
apply (succ_int IHx). Definition Z'_to_int_pos_homomorphism n :
Z'_to_Int (positive (S n)) = succ_int (Z'_to_Int (positive n)) := idpath.
Focus 2.
intro x. Definition Z'_to_int_neg_homomorphism n :
induction x. Z'_to_Int (negative (S n)) = pred_int (Z'_to_Int (negative n)) := idpath.
apply (Int.zero).
apply (pred_int IHx). Theorem Int_to_Z'_to_Int : forall x : Int, Z'_to_Int(Int_to_Z' x) = x.
Proof.
simpl. induction x as [p | | p].
rewrite transport_const. - induction p as [ | p IHp ].
reflexivity. * reflexivity.
Defined. * refine (Z'_to_int_neg_homomorphism _ @ ap pred_int IHp).
- reflexivity.
Definition Pos_to_Nat : Pos -> nat. - induction p as [ | p IHp ].
Proof. * reflexivity.
intro x. * refine (Z'_to_int_pos_homomorphism _ @ ap succ_int IHp).
induction x. Defined.
apply 1.
apply (S IHx). Lemma Int_to_Z'_succ_homomorphism :
Defined. forall x, Int_to_Z' (succ_int x) = succ_Z' (Int_to_Z' x).
Proof.
induction x as [p | | p].
Definition Int_to_Z' (x : Int) : Z'. - induction p as [ | p IHp] ; cbn.
Proof. * apply path.
induction x. * reflexivity.
apply (negative (Pos_to_Nat p)). - reflexivity.
apply (positive 0). - induction p ; reflexivity.
apply (positive (Pos_to_Nat p)). Defined.
Defined.
Lemma Int_to_Z'_pred_homomorphism :
Lemma Z'_to_int_pos_homomorphism : forall x : Int, Int_to_Z' (pred_int x) = pred_Z' (Int_to_Z' x).
forall n : nat, Z'_to_Int (positive (S n)) = succ_int (Z'_to_Int (positive n)). Proof.
Proof. induction x as [p | | p] ; try (induction p) ; reflexivity.
intro n. Defined.
reflexivity.
Qed. Theorem Z'_to_Int_to_Z' : forall x : Z', Int_to_Z'(Z'_to_Int x) = x.
Proof.
Lemma Z'_to_int_neg_homomorphism : simple refine (Z'_ind _ _ _ _) ; simpl.
forall n : nat, Z'_to_Int (negative (S n)) = pred_int (Z'_to_Int (negative n)). - induction x as [ | x] ; simpl.
Proof. * reflexivity.
intro n. * refine (ap Int_to_Z' (Z'_to_int_pos_homomorphism _)^ @ _) ; simpl.
reflexivity. refine (Int_to_Z'_succ_homomorphism _ @ _).
Qed. apply (ap succ_Z' IHx).
- induction x as [ | x IHx] ; simpl.
Theorem isoEq1 : forall x : Int, Z'_to_Int(Int_to_Z' x) = x. * apply path.
Proof. * refine (ap Int_to_Z' (Z'_to_int_neg_homomorphism _)^ @ _) ; simpl.
intro x. refine (Int_to_Z'_pred_homomorphism _ @ _).
induction x. apply (ap pred_Z' IHx).
induction p. - simpl.
reflexivity. refine (transport_paths_FlFr path _ @ _).
refine (ap (fun z => _ @ z) (ap_idmap _) @ _).
rewrite Z'_to_int_neg_homomorphism. refine (ap (fun z => z @ _) (concat_p1 _) @ _).
rewrite IHp. assert (ap (fun x : Z' => Z'_to_Int x) path = idpath) as X.
reflexivity. {
apply axiomK_hset ; apply hset_int.
reflexivity. }
refine (ap (fun z => z^ @ path) (ap_compose Z'_to_Int Int_to_Z' path) @ _).
induction p. refine (ap (fun z => (ap Int_to_Z' z)^ @ _) X @ _).
reflexivity. apply concat_1p.
Defined.
rewrite Z'_to_int_pos_homomorphism.
rewrite IHp. Definition biinv_Z'_to_Int : BiInv Z'_to_Int :=
reflexivity. ((Int_to_Z' ; Z'_to_Int_to_Z'), (Int_to_Z' ; Int_to_Z'_to_Int)).
Defined.
Definition equiv_Z'_to_Int : IsEquiv Z'_to_Int :=
Lemma Int_to_Z'_succ_homomorphism : isequiv_biinv _ biinv_Z'_to_Int.
forall x : Int, Int_to_Z' (succ_int x) = succ_Z' (Int_to_Z' x).
Proof. Instance Z'_set : IsHSet Z'.
simpl. Proof.
intro x. intros x y p q.
simpl.
induction x. Admitted.
induction p.
compute. Definition Z_to_Z' : Z -> Z'.
apply path. Proof.
hrecursion.
reflexivity. - apply (positive 0).
- apply succ_Z'.
reflexivity. - apply pred_Z'.
- hinduction.
induction p. * apply (fun _ => idpath).
reflexivity. * induction x as [ | x IHx] ; simpl.
** apply path^.
reflexivity. ** reflexivity.
Qed. * apply set_path2.
- hinduction.
Lemma Int_to_Z'_pred_homomorphism : * induction x as [ | x IHx] ; simpl.
forall x : Int, Int_to_Z' (pred_int x) = pred_Z' (Int_to_Z' x). ** apply path.
Proof. ** reflexivity.
intro x. * apply (fun _ => idpath).
induction x. * apply set_path2.
induction p. Defined.
reflexivity.
Definition Z'_to_Z : Z' -> Z.
reflexivity. Proof.
hrecursion.
reflexivity. - induction 1 as [ | x IHx].
* apply zero_Z.
induction p. * apply (succ IHx).
reflexivity. - induction 1 as [ | x IHx].
* apply zero_Z.
reflexivity. * apply (pred IHx).
- reflexivity.
Qed. Defined.
Theorem isoEq2 : forall x : Z', Int_to_Z'(Z'_to_Int x) = x. Theorem Z'_to_Z_to_Z' : forall n : Z', Z_to_Z'(Z'_to_Z n) = n.
Proof. Proof.
refine (Z'_ind _ _ _ _). hinduction.
Unshelve. - induction x as [ | x IHx] ; cbn.
* reflexivity.
Focus 2. * apply (ap succ_Z' IHx).
intro x. - induction x as [ | x IHx] ; cbn.
induction x. * apply path.
reflexivity. * apply (ap pred_Z' IHx).
rewrite Z'_to_int_pos_homomorphism. - apply set_path2.
rewrite Int_to_Z'_succ_homomorphism. Defined.
rewrite IHx.
reflexivity. Lemma Z'_to_Z_succ : forall n, Z'_to_Z(succ_Z' n) = succ(Z'_to_Z n).
Proof.
Focus 2. hinduction.
intro x. - apply (fun _ => idpath).
induction x. - induction x.
apply path. * reflexivity.
rewrite Z'_to_int_neg_homomorphism. * apply inv2.
rewrite Int_to_Z'_pred_homomorphism. - apply set_path2.
rewrite IHx. Defined.
reflexivity.
Lemma Z'_to_Z_pred : forall n, Z'_to_Z(pred_Z' n) = pred(Z'_to_Z n).
simpl. Proof.
rewrite @HoTT.Types.Paths.transport_paths_FlFr. hinduction.
rewrite concat_p1. - induction x.
rewrite ap_idmap. * reflexivity.
* apply inv1.
enough (ap (fun x : Z' => Z'_to_Int x) path = reflexivity Int.zero). - apply (fun _ => idpath).
rewrite ap_compose. - apply set_path2.
rewrite X. Defined.
apply concat_1p.
Theorem Z_to_Z'_to_Z : forall n : Z, Z'_to_Z(Z_to_Z' n) = n.
apply axiomK_hset. Proof.
apply hset_int. hinduction ; try (intros ; apply set_path2).
Defined. - reflexivity.
- intros x Hx.
Theorem adj : refine (_ @ ap succ Hx).
forall x : Z', isoEq1 (Z'_to_Int x) = ap Z'_to_Int (isoEq2 x). apply Z'_to_Z_succ.
Proof. - intros x Hx.
intro x. refine (_ @ ap pred Hx).
apply hset_int. apply Z'_to_Z_pred.
Defined. Defined.
Definition isomorphism : IsEquiv Z'_to_Int. Definition biinv_Z'_to_Z : BiInv Z'_to_Z :=
Proof. ((Z_to_Z' ; Z'_to_Z_to_Z'), (Z_to_Z' ; Z_to_Z'_to_Z)).
apply (BuildIsEquiv Z' Int Z'_to_Int Int_to_Z' isoEq1 isoEq2 adj).
Qed. Definition equiv_Z'_to_Z : IsEquiv Z'_to_Z :=
isequiv_biinv _ biinv_Z'_to_Z.
Axiom everythingSet : forall T : Type, IsHSet T.
End Isomorphic.
Definition Z_to_Z' : Z -> Z'.
Proof.
refine (Z_rec _ _ _ _ _ _).
Unshelve.
Focus 1.
apply (positive 0).
Focus 3.
apply succ_Z'.
Focus 3.
apply pred_Z'.
Focus 1.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 2.
intros.
reflexivity.
Focus 2.
intros.
induction x.
Focus 1.
compute.
apply path^.
reflexivity.
apply everythingSet.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 2.
intros.
induction x.
Focus 1.
compute.
apply path.
reflexivity.
Focus 2.
intros.
reflexivity.
apply everythingSet.
Defined.
Definition Z'_to_Z : Z' -> Z.
Proof.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 2.
induction 1.
apply nul.
apply (succ IHx).
Focus 2.
induction 1.
Focus 1.
apply nul.
apply (pred IHx).
simpl.
rewrite transport_const.
reflexivity.
Defined.
Theorem isoZEq1 : forall n : Z', Z_to_Z'(Z'_to_Z n) = n.
Proof.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 3.
intros.
induction x.
compute.
apply path.
transitivity (Z_to_Z' (pred (Z'_to_Z (negative x)))).
enough (Z'_to_Z (negative x.+1) = pred (Z'_to_Z (negative x))).
rewrite X.
reflexivity.
reflexivity.
transitivity (pred_Z' (Z_to_Z' (Z'_to_Z (negative x)))).
Focus 1.
reflexivity.
rewrite IHx.
reflexivity.
Focus 2.
intros.
induction x.
Focus 1.
reflexivity.
transitivity (Z_to_Z' (succ (Z'_to_Z (positive x)))).
reflexivity.
transitivity (succ_Z' (Z_to_Z' (Z'_to_Z (positive x)))).
reflexivity.
rewrite IHx.
reflexivity.
apply everythingSet.
Defined.
Theorem isoZEq2 : forall n : Z, Z'_to_Z(Z_to_Z' n) = n.
Proof.
refine (Z_ind _ _ _ _ _ _).
Unshelve.
Focus 1.
reflexivity.
Focus 1.
intros.
apply everythingSet.
Focus 1.
intros.
apply everythingSet.
Focus 1.
intro n.
intro X.
transitivity (Z'_to_Z (succ_Z' (Z_to_Z' n))).
reflexivity.
transitivity (succ (Z'_to_Z (Z_to_Z' n))).
Focus 2.
rewrite X.
reflexivity.
enough (forall m : Z', Z'_to_Z (succ_Z' m) = succ (Z'_to_Z m)).
rewrite X0.
reflexivity.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 2.
intros.
reflexivity.
Focus 2.
intros.
induction x.
Focus 1.
reflexivity.
compute.
rewrite <- inv2.
reflexivity.
apply everythingSet.
intros.
transitivity (Z'_to_Z (pred_Z' (Z_to_Z' n))).
reflexivity.
transitivity (pred (Z'_to_Z (Z_to_Z' n))).
Focus 2.
rewrite X.
reflexivity.
enough (forall m : Z', Z'_to_Z (pred_Z' m) = pred (Z'_to_Z m)).
rewrite X0.
reflexivity.
refine (Z'_ind _ _ _ _).
Unshelve.
Focus 1.
apply everythingSet.
Focus 1.
intro x.
induction x.
reflexivity.
compute.
rewrite <- inv1.
reflexivity.
intro x.
reflexivity.
Defined.
Theorem adj2 :
forall x : Z', isoZEq2 (Z'_to_Z x) = ap Z'_to_Z (isoZEq1 x).
Proof.
intro x.
apply everythingSet.
Defined.
Definition isomorphism2 : IsEquiv Z'_to_Z.
Proof.
apply (BuildIsEquiv Z' Z Z'_to_Z Z_to_Z' isoZEq2 isoZEq1 adj2).
Qed.
*)