diff --git a/FinSets.v b/FinSets.v index 6bbeafc..5bab981 100644 --- a/FinSets.v +++ b/FinSets.v @@ -45,10 +45,6 @@ Fixpoint FinSets_rec (FinSets_rec A P e l u assocP commP nlP nrP idemP z) end) assocP commP nlP nrP idemP. -Definition FinSetsCL A : HitRec.class (FinSets A) _ := - HitRec.Class (FinSets A) (fun x P e l u aP cP lP rP iP => FinSets_rec A P e l u aP cP lP rP iP x). -Canonical Structure FinSetsTy A : HitRec.type := HitRec.Pack _ _ (FinSetsCL A). - Axiom FinSets_beta_assoc : forall (A : Type) (P : Type) @@ -135,6 +131,12 @@ Axiom FinSets_beta_idem : forall ap (FinSets_rec A P e l u assocP commP nlP nrP idemP) (idem A x) = idemP x. + +(* TODO: add an induction principle *) +Definition FinSetsCL A : HitRec.class (FinSets A) _ _ := + HitRec.Class (FinSets A) (fun x P e l u aP cP lP rP iP => FinSets_rec A P e l u aP cP lP rP iP x) (fun x P e l u aP cP lP rP iP => FinSets_rec A P e l u aP cP lP rP iP x). +Canonical Structure FinSetsTy A : HitRec.type := HitRec.Pack _ _ _ (FinSetsCL A). + End FinSet. Section functions. diff --git a/HitTactics.v b/HitTactics.v index d43f35b..19e082f 100644 --- a/HitTactics.v +++ b/HitTactics.v @@ -1,45 +1,29 @@ Module HitRec. Record class (H : Type) (* The HIT type itself *) + (indTy : H -> Type) (* The type of the induciton principle *) (recTy : H -> Type) (* The type of the recursion principle *) - := Class { recursion_term : forall (x : H), recTy x }. -Structure type := Pack { obj : Type ; rec : obj -> Type ; class_of : class obj rec }. + := Class { + H_recursor : forall (x : H), recTy x; + H_inductor : forall (x : H), indTy x }. +Structure type := Pack { obj : Type ; ind : obj -> Type ; rec : obj -> Type ; class_of : class obj ind rec }. Definition hrecursion (e : type) : forall x, rec e x := - let 'Pack _ _ (Class r) := e + let 'Pack _ _ _ (Class r _) := e in r. +Definition hinduction (e : type) : forall x, ind e x := + let 'Pack _ _ _ (Class _ i) := e + in i. Arguments hrecursion {e} x : simpl never. -Arguments Class H {recTy} recursion_term. +Arguments hinduction {e} x : simpl never. +Arguments Class H {indTy recTy} H_recursor H_inductor. End HitRec. Ltac hrecursion_ target := generalize dependent target; match goal with - | [ |- forall target, ?P] => intro target; + | [ |- _ -> ?P ] => intro target; match type of (HitRec.hrecursion target) with | ?Q => match (eval simpl in Q) with - | forall Y , Y target => - simple refine (HitRec.hrecursion target (fun target => P) _) - | forall Y _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _) - | forall Y _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _) - | forall Y _ _ _, Y target => - let hrec:=(eval hnf in (HitRec.hrecursion target)) in - simple refine (hrec (fun target => P) _ _ _) - | forall Y _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _) - | forall Y _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _) - | forall Y _ _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _ _) - | forall Y _ _ _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _ _ _) - | forall Y _ _ _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _ _ _) - | forall Y _ _ _ _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _ _ _ _) - | forall Y _ _ _ _ _ _ _ _ _, Y target => - simple refine (HitRec.hrecursion target (fun target => P) _ _ _ _ _ _ _ _ _) | forall Y, Y => simple refine (HitRec.hrecursion target P) | forall Y _, Y => @@ -63,10 +47,40 @@ Ltac hrecursion_ target := | forall Y _ _ _ _ _ _ _ _ _ _, Y => simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _ _ _) | _ => fail "Cannot handle the recursion principle (too many parameters?) :(" + end + end + | [ |- forall target, ?P] => intro target; + match type of (HitRec.hinduction target) with + | ?Q => + match (eval simpl in Q) with + | forall Y , Y target => + simple refine (HitRec.hinduction target (fun target => P) _) + | forall Y _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _) + | forall Y _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _) + | forall Y _ _ _, Y target => + let hrec:=(eval hnf in (HitRec.hinduction target)) in + simple refine (hrec (fun target => P) _ _ _) + | forall Y _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _) + | forall Y _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _) + | forall Y _ _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _) + | forall Y _ _ _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _) + | forall Y _ _ _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _) + | forall Y _ _ _ _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _ _) + | forall Y _ _ _ _ _ _ _ _ _, Y target => + simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _ _ _) + | _ => fail "Cannot handle the induction principle (too many parameters?) :(" end end + (*| _ => fail "I am sorry, but something went wrong!"*) end. Ltac hrecursion x := hrecursion_ x; simpl; try (clear x). - - +Ltac hinduction x := hrecursion x. diff --git a/Mod2.v b/Mod2.v index 2f9bc8d..c342950 100644 --- a/Mod2.v +++ b/Mod2.v @@ -52,9 +52,9 @@ Axiom Mod2_rec_beta_mod : forall , ap (Mod2_rec P a s mod') mod = mod'. -Definition Mod2CL : HitRec.class Mod2 _ := - HitRec.Class Mod2 (fun x P a s p => Mod2_ind P a s p x). -Canonical Structure Mod2ty : HitRec.type := HitRec.Pack Mod2 _ Mod2CL. +Definition Mod2CL : HitRec.class Mod2 _ _ := + HitRec.Class Mod2 (fun x P a s p => Mod2_rec P a s p x) (fun x P a s p => Mod2_ind P a s p x). +Canonical Structure Mod2ty : HitRec.type := HitRec.Pack Mod2 _ _ Mod2CL. End modulo. @@ -79,132 +79,72 @@ intro x. hrecursion x. - apply Z. - intros. apply (succ H). -- simpl. - etransitivity. - eapply transport_const. - eapply modulo2. +- simpl. apply mod. Defined. Definition plus : Mod2 -> Mod2 -> Mod2. Proof. -intro n. -refine (Mod2_ind _ _ _ _). - Unshelve. - - Focus 2. - apply n. - - Focus 2. - intro m. - intro k. - apply (succ k). - - simpl. - rewrite transport_const. - apply modulo2. +intros n m. +hrecursion m. +- exact n. +- apply succ. +- apply modulo2. Defined. Definition Bool_to_Mod2 : Bool -> Mod2. Proof. intro b. destruct b. - apply (succ Z). - - apply Z. ++ apply (succ Z). ++ apply Z. Defined. Definition Mod2_to_Bool : Mod2 -> Bool. Proof. -refine (Mod2_ind _ _ _ _). - Unshelve. - Focus 2. - apply false. - - Focus 2. - intro n. - apply negb. - - Focus 1. - simpl. - apply transport_const. +intro x. +hrecursion x. +- exact false. +- exact negb. +- simpl. reflexivity. Defined. Theorem eq1 : forall n : Bool, Mod2_to_Bool (Bool_to_Mod2 n) = n. Proof. intro b. -destruct b. - Focus 1. - compute. - reflexivity. - - compute. - reflexivity. +destruct b; compute; reflexivity. Qed. Theorem Bool_to_Mod2_negb : forall x : Bool, succ (Bool_to_Mod2 x) = Bool_to_Mod2 (negb x). Proof. intros. -destruct x. - compute. - apply mod^. - - compute. - apply reflexivity. +destruct x; compute. ++ apply mod^. ++ apply reflexivity. Defined. Theorem eq2 : forall n : Mod2, Bool_to_Mod2 (Mod2_to_Bool n) = n. Proof. -refine (Mod2_ind _ _ _ _). - Unshelve. - Focus 2. - compute. - reflexivity. - - Focus 2. - intro n. - intro IHn. - symmetry. - transitivity (succ (Bool_to_Mod2 (Mod2_to_Bool n))). - - Focus 1. - symmetry. - apply (ap succ IHn). - - transitivity (Bool_to_Mod2 (negb (Mod2_to_Bool n))). - apply Bool_to_Mod2_negb. - enough (negb (Mod2_to_Bool n) = Mod2_to_Bool (succ n)). - apply (ap Bool_to_Mod2 X). - - compute. - reflexivity. - simpl. - rewrite concat_p1. - rewrite concat_1p. - rewrite @HoTT.Types.Paths.transport_paths_FlFr. - rewrite concat_p1. - rewrite ap_idmap. - rewrite ap_compose. - - enough (ap Mod2_to_Bool mod = reflexivity false). - rewrite X. - simpl. - rewrite concat_1p. - rewrite inv_V. - reflexivity. - - enough (IsHSet Bool). - apply axiomK_hset. - apply X. - apply hset_bool. +intro n. +hinduction n. +- reflexivity. +- intros n IHn. + symmetry. etransitivity. apply (ap succ IHn^). + etransitivity. apply Bool_to_Mod2_negb. + hott_simpl. +- rewrite @HoTT.Types.Paths.transport_paths_FlFr. + hott_simpl. + rewrite ap_compose. + enough (ap Mod2_to_Bool mod = idpath). + + rewrite X. hott_simpl. + + unfold Mod2_to_Bool. unfold HitRec.hrecursion. simpl. + apply (Mod2_rec_beta_mod Bool false negb 1). Defined. Theorem adj : forall x : Mod2, eq1 (Mod2_to_Bool x) = ap Mod2_to_Bool (eq2 x). Proof. intro x. -enough (IsHSet Bool). -apply set_path2. apply hset_bool. Defined.