Replace the canonical structures mechanism for recursion with type classes

See changes in HitTactics.v
This commit is contained in:
Dan Frumin 2017-05-22 21:06:13 +02:00
parent 5e594d10f0
commit f23d4aeacb
3 changed files with 131 additions and 169 deletions

139
FinSets.v
View File

@ -319,10 +319,9 @@ Axiom FSet_ind_beta_idem : forall
End FSet. End FSet.
(* TODO: add an induction principle *) Instance FSet_recursion A : HitRecursion (FSet A) := {
Definition FSetCL A : HitRec.class (FSet A) _ _ := indTy := _; recTy := _;
HitRec.Class (FSet A) (fun x P H e l u aP cP lP rP iP => FSet_rec A P H e l u aP cP lP rP iP x) (fun x P H e l u aP cP lP rP iP => FSet_ind A P H e l u aP cP lP rP iP x). H_inductor := FSet_ind A; H_recursor := FSet_rec A }.
Canonical Structure FSetTy A : HitRec.type := HitRec.Pack _ _ _ (FSetCL A).
Arguments E {_}. Arguments E {_}.
Arguments U {_} _ _. Arguments U {_} _ _.
@ -332,28 +331,45 @@ End FinSet.
Section functions. Section functions.
Parameter A : Type. Parameter A : Type.
Parameter eq : A -> A -> Bool. Parameter A_eqdec : forall (x y : A), Decidable (x = y).
Parameter eq_refl: forall a:A, eq a a = true. Definition deceq (x y : A) :=
if dec (x = y) then true else false.
Definition isIn : A -> FSet A -> Bool. Definition isIn : A -> FSet A -> Bool.
Proof. Proof.
intros a X. intros a.
hrecursion X. hrecursion.
- exact false. - exact false.
- intro a'. apply (eq a a'). - intro a'. apply (deceq a a').
- apply orb. - apply orb.
- intros x y z. compute. destruct x; reflexivity. - intros x y z. compute. destruct x; reflexivity.
- intros x y. compute. destruct x, y; reflexivity. - intros x y. compute. destruct x, y; reflexivity.
- intros x. compute. reflexivity. - intros x. compute. reflexivity.
- intros x. compute. destruct x; reflexivity. - intros x. compute. destruct x; reflexivity.
- intros a'. compute. destruct (eq a a'); reflexivity. - intros a'. compute. destruct (A_eqdec a a'); reflexivity.
Defined. Defined.
Lemma isIn_eq a b : isIn a (L b) = true -> a = b.
Proof. cbv.
destruct (A_eqdec a b). intro. apply p.
intro X.
pose (false_ne_true X). contradiction.
Defined.
Lemma isIn_empt_false a : isIn a E = true -> Empty.
Proof.
cbv. intro X.
pose (false_ne_true X). contradiction.
Defined.
Lemma isIn_union a X Y : isIn a (U X Y) = orb (isIn a X) (isIn a Y).
Proof. reflexivity. Qed.
Definition comprehension : Definition comprehension :
(A -> Bool) -> FSet A -> FSet A. (A -> Bool) -> FSet A -> FSet A.
Proof. Proof.
intros P X. intros P.
hrecursion X. hrecursion.
- apply E. - apply E.
- intro a. - intro a.
refine (if (P a) then L a else E). refine (if (P a) then L a else E).
@ -367,29 +383,40 @@ hrecursion X.
+ apply idem. + apply idem.
+ apply nl. + apply nl.
Defined. Defined.
Lemma comprehension_false Y : comprehension (fun a => isIn a E) Y = E.
Proof.
hrecursion Y; try (intros; apply set_path2).
- cbn. reflexivity.
- cbn. reflexivity.
- intros x y IHa IHb.
cbn.
rewrite IHa.
rewrite IHb.
rewrite nl.
reflexivity.
Defined.
Require Import FunextAxiom. Require Import FunextAxiom.
Lemma comprehension_idem: Lemma comprehension_idem:
forall (X:FSet A), forall Y, comprehension (fun x => isIn x (U X Y)) X = X. forall (X:FSet A), forall Y, comprehension (fun x => isIn x (U X Y)) X = X.
Proof. Proof.
simple refine (FSet_ind _ _ _ _ _ _ _ _ _ _ _); simpl. hinduction; try (intros; apply set_path2).
- intro Y. cbv. reflexivity. - intro Y. cbv. reflexivity.
- intros a Y. unfold comprehension. unfold HitRec.hrecursion. simpl. - intros a Y. cbn.
enough (isIn a (U (L a) Y) = true). unfold deceq;
+ rewrite X. reflexivity. destruct (dec (a = a)); simpl.
+ unfold isIn. unfold HitRec.hrecursion. simpl. + reflexivity.
rewrite eq_refl. auto. + contradiction n. reflexivity.
- intros X1 X2 IH1 IH2 Y. - intros X1 X2 IH1 IH2 Y.
unfold comprehension. unfold HitRec.hrecursion. simpl. cbn -[isIn].
rewrite <- (assoc _ X1 X2 Y).
f_ap. f_ap.
+ apply (IH1 (U X2 Y)). + rewrite <- assoc. apply (IH1 (U X2 Y)).
+ rewrite (assoc _ X1 X2 Y). + rewrite (comm _ X1 X2).
rewrite (comm _ X1 X2).
rewrite <- (assoc _ X2 X1 Y). rewrite <- (assoc _ X2 X1 Y).
apply (IH2 (U X1 Y)). apply (IH2 (U X1 Y)).
Admitted. Admitted.
Definition intersection : Definition intersection :
FSet A -> FSet A -> FSet A. FSet A -> FSet A -> FSet A.
Proof. Proof.
@ -411,68 +438,6 @@ hrecursion x.
destruct (isIn a y); reflexivity. destruct (isIn a y); reflexivity.
Defined. Defined.
Definition subset' (x : FSet A) (y : FSet A) : Bool.
Proof.
refine (FSet_rec A _ _ _ _ _ _ _ _ _ _).
Unshelve.
Focus 6.
apply x.
Focus 6.
apply true.
Focus 6.
intro a.
apply (isIn a y).
Focus 6.
intro b.
intro b'.
apply (andb b b').
Focus 1.
intros.
compute.
destruct x0.
destruct y0.
reflexivity.
reflexivity.
reflexivity.
Focus 1.
intros.
compute.
destruct x0.
destruct y0.
reflexivity.
reflexivity.
destruct y0.
reflexivity.
reflexivity.
Focus 1.
intros.
compute.
reflexivity.
Focus 1.
intros.
compute.
destruct x0.
reflexivity.
reflexivity.
intros.
destruct (isIn x0 y).
compute.
reflexivity.
compute.
reflexivity.
Defined.
(* TODO: subset = subset' *)
Definition equal_set (x : FSet A) (y : FSet A) : Bool Definition equal_set (x : FSet A) (y : FSet A) : Bool
:= andb (subset x y) (subset y x). := andb (subset x y) (subset y x).

View File

@ -1,86 +1,85 @@
Module HitRec. Class HitRecursion (H : Type) := {
Record class (H : Type) (* The HIT type itself *) indTy : Type;
(indTy : H -> Type) (* The type of the induciton principle *) recTy : Type;
(recTy : H -> Type) (* The type of the recursion principle *) H_inductor : indTy;
:= Class { H_recursor : recTy
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
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 hinduction {e} x : simpl never.
Arguments Class H {indTy recTy} H_recursor H_inductor.
End HitRec.
Ltac hrecursion_ target := Definition hrecursion (H : Type) {HR : HitRecursion H} : @recTy H HR :=
generalize dependent target; @H_recursor H HR.
match goal with
| [ |- _ -> ?P ] => intro target; Definition hinduction (H : Type) {HR : HitRecursion H} : @indTy H HR :=
match type of (HitRec.hrecursion target) with @H_inductor H HR.
Ltac hrecursion_ :=
lazymatch goal with
| [ |- ?T -> ?P ] =>
let hrec1 := eval cbv[hrecursion H_recursor] in (hrecursion T) in
let hrec := eval simpl in hrec1 in
match type of hrec with
| ?Q => | ?Q =>
match (eval simpl in Q) with match (eval simpl in Q) with
| forall Y, Y => | forall Y, T -> Y =>
simple refine (HitRec.hrecursion target P) simple refine (hrec P)
| forall Y _, Y => | forall Y _, T -> Y =>
simple refine (HitRec.hrecursion target P _) simple refine (hrec P _)
| forall Y _ _, Y => | forall Y _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _) simple refine (hrec P _ _)
| forall Y _ _ _, Y => | forall Y _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _) simple refine (hrec P _ _ _)
| forall Y _ _ _ _, Y => | forall Y _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _) simple refine (hrec P _ _ _ _)
| forall Y _ _ _ _ _, Y => | forall Y _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _) simple refine (hrec P _ _ _ _ _)
| forall Y _ _ _ _ _ _, Y => | forall Y _ _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _) simple refine (hrec P _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _, Y => | forall Y _ _ _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _) simple refine (hrec P _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _, Y => | forall Y _ _ _ _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _) simple refine (hrec P _ _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _ _, Y => | forall Y _ _ _ _ _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _ _) simple refine (hrec P _ _ _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _ _ _, Y => | forall Y _ _ _ _ _ _ _ _ _ _, T -> Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _ _ _) simple refine (hrec P _ _ _ _ _ _ _ _ _ _)
| _ => 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; | [ |- forall (target:?T), ?P] =>
match type of (HitRec.hinduction target) with let hind1 := eval cbv[hinduction H_inductor] in (hinduction T) in
let hind := eval simpl in hind1 in
match type of hind with
| ?Q => | ?Q =>
match (eval simpl in Q) with match (eval simpl in Q) with
| forall Y , Y target => | forall Y , (forall x, Y x) =>
simple refine (HitRec.hinduction target (fun target => P) _) simple refine (hind (fun target => P) _)
| forall Y _, Y target => | forall Y _, (forall x, Y x) =>
simple refine (HitRec.hinduction target (fun target => P) _) simple refine (hind (fun target => P) _)
| forall Y _ _, Y target => | forall Y _ _, (forall x, Y x) =>
simple refine (HitRec.hinduction target (fun target => P) _ _) simple refine (hind (fun target => P) _ _)
| forall Y _ _ _, Y target => | forall Y _ _ _, (forall x, Y x) =>
let hrec:=(eval hnf in (HitRec.hinduction target)) in simple refine (hind (fun target => P) _ _ _)
simple refine (hrec (fun target => P) _ _ _) | forall Y _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _) | forall Y _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _) | forall Y _ _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _) | forall Y _ _ _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _) | forall Y _ _ _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _) | forall Y _ _ _ _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _ _) | forall Y _ _ _ _ _ _ _ _ _, (forall x, Y x) =>
| forall Y _ _ _ _ _ _ _ _ _, Y target => simple refine (hind (fun target => P) _ _ _ _ _ _ _ _ _)
simple refine (HitRec.hinduction target (fun target => P) _ _ _ _ _ _ _ _ _)
| _ => fail "Cannot handle the induction principle (too many parameters?) :(" | _ => fail "Cannot handle the induction principle (too many parameters?) :("
end end
end end
(*| _ => fail "I am sorry, but something went wrong!"*) (*| _ => fail "I am sorry, but something went wrong!"*)
end. end.
Ltac hrecursion x := hrecursion_ x; simpl; try (clear x). Tactic Notation "hrecursion" := hrecursion_; simpl.
Ltac hinduction x := hrecursion x. Tactic Notation "hrecursion" ident(x) := revert x; hrecursion.
Tactic Notation "hinduction" := hrecursion_; simpl.
Tactic Notation "hinduction" ident(x) := revert x; hrecursion.

18
Mod2.v
View File

@ -1,6 +1,6 @@
Require Export HoTT. Require Export HoTT.
Require Import HitTactics. Require Import HitTactics.
Module Export modulo. Module Export modulo.
Private Inductive Mod2 : Type0 := Private Inductive Mod2 : Type0 :=
@ -51,10 +51,10 @@ Axiom Mod2_rec_beta_mod : forall
(mod' : a = s (s a)) (mod' : a = s (s a))
, ap (Mod2_rec P a s mod') mod = mod'. , ap (Mod2_rec P a s mod') mod = mod'.
Instance: HitRecursion Mod2 := {
Definition Mod2CL : HitRec.class Mod2 _ _ := indTy := _; recTy := _;
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). H_inductor := Mod2_ind;
Canonical Structure Mod2ty : HitRec.type := HitRec.Pack Mod2 _ _ Mod2CL. H_recursor := Mod2_rec }.
End modulo. End modulo.
@ -62,7 +62,7 @@ End modulo.
Theorem modulo2 : forall n : Mod2, n = succ(succ n). Theorem modulo2 : forall n : Mod2, n = succ(succ n).
Proof. Proof.
intro n. intro n.
hrecursion n. hinduction n.
- apply mod. - apply mod.
- intros n p. - intros n p.
apply (ap succ p). apply (ap succ p).
@ -75,8 +75,7 @@ Defined.
Definition negate : Mod2 -> Mod2. Definition negate : Mod2 -> Mod2.
Proof. Proof.
intro x. hrecursion.
hrecursion x.
- apply Z. - apply Z.
- intros. apply (succ H). - intros. apply (succ H).
- simpl. apply mod. - simpl. apply mod.
@ -137,8 +136,7 @@ hinduction n.
rewrite ap_compose. rewrite ap_compose.
enough (ap Mod2_to_Bool mod = idpath). enough (ap Mod2_to_Bool mod = idpath).
+ rewrite X. hott_simpl. + rewrite X. hott_simpl.
+ unfold Mod2_to_Bool. unfold HitRec.hrecursion. simpl. + apply (Mod2_rec_beta_mod Bool false negb 1).
apply (Mod2_rec_beta_mod Bool false negb 1).
Defined. Defined.
Theorem adj : Theorem adj :