Package both the induction and recursion principles in the same record

This commit is contained in:
Dan Frumin 2017-05-20 14:05:46 +02:00
parent 44da34d72f
commit 399fab467b
3 changed files with 84 additions and 128 deletions

View File

@ -45,10 +45,6 @@ Fixpoint FinSets_rec
(FinSets_rec A P e l u assocP commP nlP nrP idemP z) (FinSets_rec A P e l u assocP commP nlP nrP idemP z)
end) assocP commP nlP nrP idemP. 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 Axiom FinSets_beta_assoc : forall
(A : Type) (A : Type)
(P : 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) ap (FinSets_rec A P e l u assocP commP nlP nrP idemP) (idem A x)
= =
idemP 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. End FinSet.
Section functions. Section functions.

View File

@ -1,45 +1,29 @@
Module HitRec. Module HitRec.
Record class (H : Type) (* The HIT type itself *) 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 *) (recTy : H -> Type) (* The type of the recursion principle *)
:= Class { recursion_term : forall (x : H), recTy x }. := Class {
Structure type := Pack { obj : Type ; rec : obj -> Type ; class_of : class obj rec }. 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 := Definition hrecursion (e : type) : forall x, rec e x :=
let 'Pack _ _ (Class r) := e let 'Pack _ _ _ (Class r _) := e
in r. 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 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. End HitRec.
Ltac hrecursion_ target := Ltac hrecursion_ target :=
generalize dependent target; generalize dependent target;
match goal with match goal with
| [ |- forall target, ?P] => intro target; | [ |- _ -> ?P ] => intro target;
match type of (HitRec.hrecursion target) with match type of (HitRec.hrecursion target) with
| ?Q => | ?Q =>
match (eval simpl in Q) with 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 => | forall Y, Y =>
simple refine (HitRec.hrecursion target P) simple refine (HitRec.hrecursion target P)
| forall Y _, Y => | forall Y _, Y =>
@ -65,8 +49,38 @@ Ltac hrecursion_ target :=
| _ => fail "Cannot handle the recursion principle (too many parameters?) :(" | _ => fail "Cannot handle the recursion principle (too many parameters?) :("
end end
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. end.
Ltac hrecursion x := hrecursion_ x; simpl; try (clear x). Ltac hrecursion x := hrecursion_ x; simpl; try (clear x).
Ltac hinduction x := hrecursion x.

126
Mod2.v
View File

@ -52,9 +52,9 @@ Axiom Mod2_rec_beta_mod : forall
, ap (Mod2_rec P a s mod') mod = mod'. , ap (Mod2_rec P a s mod') mod = mod'.
Definition Mod2CL : HitRec.class Mod2 _ := Definition Mod2CL : HitRec.class Mod2 _ _ :=
HitRec.Class Mod2 (fun x P a s p => Mod2_ind P a s p x). 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. Canonical Structure Mod2ty : HitRec.type := HitRec.Pack Mod2 _ _ Mod2CL.
End modulo. End modulo.
@ -79,132 +79,72 @@ intro x.
hrecursion x. hrecursion x.
- apply Z. - apply Z.
- intros. apply (succ H). - intros. apply (succ H).
- simpl. - simpl. apply mod.
etransitivity.
eapply transport_const.
eapply modulo2.
Defined. Defined.
Definition plus : Mod2 -> Mod2 -> Mod2. Definition plus : Mod2 -> Mod2 -> Mod2.
Proof. Proof.
intro n. intros n m.
refine (Mod2_ind _ _ _ _). hrecursion m.
Unshelve. - exact n.
- apply succ.
Focus 2. - apply modulo2.
apply n.
Focus 2.
intro m.
intro k.
apply (succ k).
simpl.
rewrite transport_const.
apply modulo2.
Defined. Defined.
Definition Bool_to_Mod2 : Bool -> Mod2. Definition Bool_to_Mod2 : Bool -> Mod2.
Proof. Proof.
intro b. intro b.
destruct b. destruct b.
apply (succ Z). + apply (succ Z).
+ apply Z.
apply Z.
Defined. Defined.
Definition Mod2_to_Bool : Mod2 -> Bool. Definition Mod2_to_Bool : Mod2 -> Bool.
Proof. Proof.
refine (Mod2_ind _ _ _ _). intro x.
Unshelve. hrecursion x.
Focus 2. - exact false.
apply false. - exact negb.
- simpl. reflexivity.
Focus 2.
intro n.
apply negb.
Focus 1.
simpl.
apply transport_const.
Defined. Defined.
Theorem eq1 : forall n : Bool, Mod2_to_Bool (Bool_to_Mod2 n) = n. Theorem eq1 : forall n : Bool, Mod2_to_Bool (Bool_to_Mod2 n) = n.
Proof. Proof.
intro b. intro b.
destruct b. destruct b; compute; reflexivity.
Focus 1.
compute.
reflexivity.
compute.
reflexivity.
Qed. Qed.
Theorem Bool_to_Mod2_negb : forall x : Bool, Theorem Bool_to_Mod2_negb : forall x : Bool,
succ (Bool_to_Mod2 x) = Bool_to_Mod2 (negb x). succ (Bool_to_Mod2 x) = Bool_to_Mod2 (negb x).
Proof. Proof.
intros. intros.
destruct x. destruct x; compute.
compute. + apply mod^.
apply mod^. + apply reflexivity.
compute.
apply reflexivity.
Defined. Defined.
Theorem eq2 : forall n : Mod2, Bool_to_Mod2 (Mod2_to_Bool n) = n. Theorem eq2 : forall n : Mod2, Bool_to_Mod2 (Mod2_to_Bool n) = n.
Proof. Proof.
refine (Mod2_ind _ _ _ _). intro n.
Unshelve. hinduction n.
Focus 2. - reflexivity.
compute. - intros n IHn.
reflexivity. symmetry. etransitivity. apply (ap succ IHn^).
etransitivity. apply Bool_to_Mod2_negb.
Focus 2. hott_simpl.
intro n. - rewrite @HoTT.Types.Paths.transport_paths_FlFr.
intro IHn. hott_simpl.
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. rewrite ap_compose.
enough (ap Mod2_to_Bool mod = idpath).
enough (ap Mod2_to_Bool mod = reflexivity false). + rewrite X. hott_simpl.
rewrite X. + unfold Mod2_to_Bool. unfold HitRec.hrecursion. simpl.
simpl. apply (Mod2_rec_beta_mod Bool false negb 1).
rewrite concat_1p.
rewrite inv_V.
reflexivity.
enough (IsHSet Bool).
apply axiomK_hset.
apply X.
apply hset_bool.
Defined. Defined.
Theorem adj : Theorem adj :
forall x : Mod2, eq1 (Mod2_to_Bool x) = ap Mod2_to_Bool (eq2 x). forall x : Mod2, eq1 (Mod2_to_Bool x) = ap Mod2_to_Bool (eq2 x).
Proof. Proof.
intro x. intro x.
enough (IsHSet Bool).
apply set_path2.
apply hset_bool. apply hset_bool.
Defined. Defined.