Add a general hit recursion tactic.

The tactic resolves the induction principle for a HIT
using Canonical Structures, following the suggestion of @gallais
This commit is contained in:
Dan Frumin 2017-05-18 17:43:19 +02:00
parent e81986a920
commit 44da34d72f
4 changed files with 162 additions and 151 deletions

180
FinSets.v
View File

@ -1,5 +1,5 @@
Require Import HoTT.
Require Export HoTT. Require Export HoTT.
Require Import HitTactics.
Module Export FinSet. Module Export FinSet.
@ -45,6 +45,10 @@ 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)
@ -133,129 +137,66 @@ Axiom FinSets_beta_idem : forall
idemP x. idemP x.
End FinSet. End FinSet.
Definition isIn : forall Section functions.
(A : Type) Parameter A : Type.
(eq : A -> A -> Bool), Parameter eq : A -> A -> Bool.
A -> FinSets A -> Bool. Definition isIn : A -> FinSets A -> Bool.
Proof. Proof.
intro A. intros a X.
intro eq. hrecursion X.
intro a. - exact false.
refine (FinSets_rec A _ _ _ _ _ _ _ _ _). - intro a'. apply (eq a a').
Unshelve. - apply orb.
- intros x y z. compute. destruct x; reflexivity.
Focus 6. - intros x y. compute. destruct x, y; reflexivity.
apply false. - intros x. compute. reflexivity.
- intros x. compute. destruct x; reflexivity.
Focus 6. - intros a'. compute. destruct (eq a a'); reflexivity.
intro a'.
apply (eq a a').
Focus 6.
intro b.
intro b'.
apply (orb b b').
Focus 3.
intros.
compute.
reflexivity.
Focus 1.
intros.
compute.
destruct x.
reflexivity.
destruct y.
reflexivity.
reflexivity.
Focus 1.
intros.
compute.
destruct x.
destruct y.
reflexivity.
reflexivity.
destruct y.
reflexivity.
reflexivity.
Focus 1.
intros.
compute.
destruct x.
reflexivity.
reflexivity.
intros.
compute.
destruct (eq a x).
reflexivity.
reflexivity.
Defined. Defined.
Definition comprehension : forall
(A : Type) Definition comprehension :
(eq : A -> A -> Bool),
(A -> Bool) -> FinSets A -> FinSets A. (A -> Bool) -> FinSets A -> FinSets A.
Proof. Proof.
intro A. intros P X.
intro eq. hrecursion X.
intro phi. - apply empty.
refine (FinSets_rec A _ _ _ _ _ _ _ _ _). - intro a.
Unshelve. refine (if (P a) then L A a else empty A).
- intros X Y.
Focus 6. apply (U A X Y).
apply empty. - intros. cbv. apply assoc.
- intros. cbv. apply comm.
Focus 6. - intros. cbv. apply nl.
intro a. - intros. cbv. apply nr.
apply (if (phi a) then L A a else (empty A)). - intros. cbv.
destruct (P x); simpl.
Focus 6. + apply idem.
intro x. + apply nl.
intro y.
apply (U A x y).
Focus 3.
intros.
compute.
apply nl.
Focus 1.
intros.
compute.
apply assoc.
Focus 1.
intros.
compute.
apply comm.
Focus 1.
intros.
compute.
apply nr.
intros.
compute.
destruct (phi x).
apply idem.
apply nl.
Defined. Defined.
Definition intersection : forall (A : Type) (eq : A -> A -> Bool), Definition intersection :
FinSets A -> FinSets A -> FinSets A. FinSets A -> FinSets A -> FinSets A.
Proof. Proof.
intro A. intros X Y.
intro eq. apply (comprehension (fun (a : A) => isIn a X) Y).
intro x.
intro y.
apply (comprehension A eq (fun (a : A) => isIn A eq a x) y).
Defined. Defined.
Definition subset (A : Type) (eq : A -> A -> Bool) (x : FinSets A) (y : FinSets A) : Bool. Definition subset (x : FinSets A) (y : FinSets A) : Bool.
Proof.
hrecursion x.
- apply true.
- intro a. apply (isIn a y).
- apply andb.
- intros a b c. compute. destruct a; reflexivity.
- intros a b. compute. destruct a, b; reflexivity.
- intros x. compute. reflexivity.
- intros x. compute. destruct x;reflexivity.
- intros a. simpl.
destruct (isIn a y); reflexivity.
Defined.
Definition subset' (x : FinSets A) (y : FinSets A) : Bool.
Proof. Proof.
refine (FinSets_rec A _ _ _ _ _ _ _ _ _ _). refine (FinSets_rec A _ _ _ _ _ _ _ _ _ _).
Unshelve. Unshelve.
@ -268,7 +209,7 @@ refine (FinSets_rec A _ _ _ _ _ _ _ _ _ _).
Focus 6. Focus 6.
intro a. intro a.
apply (isIn A eq a y). apply (isIn a y).
Focus 6. Focus 6.
intro b. intro b.
@ -308,15 +249,16 @@ refine (FinSets_rec A _ _ _ _ _ _ _ _ _ _).
reflexivity. reflexivity.
intros. intros.
destruct (isIn A eq x0 y). destruct (isIn x0 y).
compute. compute.
reflexivity. reflexivity.
compute. compute.
reflexivity. reflexivity.
Defined. Defined.
(* TODO: subset = subset' *)
Definition equal_set (A : Type) (eq : A -> A -> Bool) (x : FinSets A) (y : FinSets A) : Bool Definition equal_set (x : FinSets A) (y : FinSets A) : Bool
:= andb (subset A eq x y) (subset A eq y x). := andb (subset x y) (subset y x).
Fixpoint eq_nat n m : Bool := Fixpoint eq_nat n m : Bool :=
match n, m with match n, m with
@ -324,4 +266,6 @@ Fixpoint eq_nat n m : Bool :=
| O, S _ => false | O, S _ => false
| S _, O => false | S _, O => false
| S n1, S m1 => eq_nat n1 m1 | S n1, S m1 => eq_nat n1 m1
end. end.
End functions.

72
HitTactics.v Normal file
View File

@ -0,0 +1,72 @@
Module HitRec.
Record class (H : Type) (* The HIT type itself *)
(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 }.
Definition hrecursion (e : type) : forall x, rec e x :=
let 'Pack _ _ (Class r) := e
in r.
Arguments hrecursion {e} x : simpl never.
Arguments Class H {recTy} recursion_term.
End HitRec.
Ltac hrecursion_ target :=
generalize dependent target;
match goal with
| [ |- forall 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 =>
simple refine (HitRec.hrecursion target P _)
| forall Y _ _, Y =>
simple refine (HitRec.hrecursion target P _ _)
| forall Y _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _)
| forall Y _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _)
| forall Y _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _)
| forall Y _ _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _ _)
| forall Y _ _ _ _ _ _ _ _ _ _, Y =>
simple refine (HitRec.hrecursion target P _ _ _ _ _ _ _ _ _ _)
| _ => fail "Cannot handle the recursion principle (too many parameters?) :("
end
end
end.
Ltac hrecursion x := hrecursion_ x; simpl; try (clear x).

60
Mod2.v
View File

@ -1,6 +1,6 @@
Require Import HoTT.
Require Export HoTT. Require Export HoTT.
Require Import HitTactics.
Module Export modulo. Module Export modulo.
Private Inductive Mod2 : Type0 := Private Inductive Mod2 : Type0 :=
@ -51,44 +51,38 @@ 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'.
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.
End modulo. End modulo.
Definition negate : Mod2 -> Mod2.
Proof.
refine (Mod2_ind _ _ _ _).
Unshelve.
Focus 2.
apply (succ Z).
Focus 2.
intros.
apply (succ H).
simpl.
rewrite transport_const.
rewrite <- mod.
reflexivity.
Defined.
Theorem modulo2 : forall n : Mod2, n = succ(succ n). Theorem modulo2 : forall n : Mod2, n = succ(succ n).
Proof. Proof.
refine (Mod2_ind _ _ _ _). intro n.
Unshelve. hrecursion n.
Focus 2. - apply mod.
apply mod. - intros n p.
apply (ap succ p).
- simpl.
etransitivity.
eapply (@transport_paths_FlFr _ _ idmap (fun n => succ (succ n))).
hott_simpl.
apply ap_compose.
Defined.
Focus 2. Definition negate : Mod2 -> Mod2.
intro n. Proof.
intro p. intro x.
apply (ap succ p). hrecursion x.
- apply Z.
simpl. - intros. apply (succ H).
rewrite @HoTT.Types.Paths.transport_paths_FlFr. - simpl.
rewrite ap_idmap. etransitivity.
rewrite concat_Vp. eapply transport_const.
rewrite concat_1p. eapply modulo2.
rewrite ap_compose.
reflexivity.
Defined. Defined.
Definition plus : Mod2 -> Mod2 -> Mod2. Definition plus : Mod2 -> Mod2 -> Mod2.

View File

@ -1,5 +1,6 @@
-R . "" COQC = hoqc COQDEP = hoqdep -R . "" COQC = hoqc COQDEP = hoqdep
HitTactics.v
Mod2.v Mod2.v
FinSets.v FinSets.v
Expressions.v Expressions.v