in
PET.mk_tactic subst_tac
-and destruct ?(simplified = false) ~first_time term =
+and destruct ~first_time term =
let are_convertible hd1 hd2 metasenv context =
fst (CR.are_convertible ~metasenv context hd1 hd2 CU.empty_ugraph)
in
| _, C.Rel i2 when DTI.does_not_occur i2 t1 ->
mk_subst_chain `RightToLeft i2 t1 t2
(* else part *)
- | _ when first_time && simplified -> raise exn_nothingtodo
- | _ when simplified (* && not first time *) -> T.id_tac
- | _ (* when not simplified *) ->
- T.then_ ~start:(simpl_in_term context term)
- ~continuation:(destruct ~simplified:true ~first_time term)
+ | _ when first_time -> raise exn_nothingtodo
+ | _ (* when not first time *) -> T.id_tac
end
- | _ when first_time && simplified -> raise exn_nothingtodo
- | _ when simplified (* && not first time *) -> T.id_tac
- | _ (* when not simplified *) ->
- T.then_ ~start:(simpl_in_term context term)
- ~continuation:(destruct ~simplified:true ~first_time term)
+ | _ when first_time -> raise exn_nothingtodo
+ | _ (* when not first time *) -> T.id_tac
in
PET.apply_tactic tac status
in
(* ORDER RELATION BETWEEN POSITIONS AND CONTEXTS
*)
-include "datatypes/Context.ma".
+include "datatypes_defs/Context.ma".
inductive CLE: Nat \to Context \to Prop \def
| cle_zero: \forall Q. CLE zero Q
*)
include "Lift/defs.ma".
-include "datatypes/Context.ma".
+include "datatypes_defs/Context.ma".
inductive Insert (p,q:Proof) (S:Sequent):
Nat \to Context \to Context \to Prop \def
(* PROOF RELOCATION
*)
-include "datatypes/Proof.ma".
+include "datatypes_defs/Proof.ma".
inductive Lift: Nat \to Nat \to Proof \to Proof \to Prop \def
| lift_lref_lt: \forall jd,jh,i. i < jd \to Lift jd jh (lref i) (lref i)
(* EQUALITY PREDICATE FOR PROOFS IN CONTEXT
*)
-include "datatypes/Context.ma".
+include "datatypes_defs/Context.ma".
inductive PEq (Q1:Context) (p1:Proof) (Q2:Context) (p2:Proof): Prop \def
| peq_intro: Q1 = Q2 \land p1 = p2 \to PEq Q1 p1 Q2 p2
(**)
+include "datatypes_props/Sequent.ma".
include "Track/inv.ma".
include "PRed/defs.ma".
lapply linear track_inv_lref to H5; decompose; autobatch
| lapply linear track_inv_scut to H3; decompose; destruct;
lapply linear track_inv_prin to H5; destruct;
- lapply linear rinj_inj to Hcut1;
- rewrite < Hcut1 in H4; clear Hcut1 a2;
- autobatch
+ lapply linear rinj_inj to Hcut1; destruct; autobatch
| lapply linear track_inv_scut to H3; decompose; destruct;
- lapply linear track_inv_prin to H4; destruct; autobatch
+ lapply linear track_inv_prin to H4; destruct;
+ lapply linear linj_inj to Hcut; destruct; autobatch
| lapply linear track_inv_scut to H3; decompose; destruct;
lapply linear track_inv_impw to H4; decompose; destruct;
lapply linear track_inv_impr to H5; decompose; destruct; autobatch
For cut elimination and confluence
*)
-include "datatypes/Context.ma".
+include "datatypes_defs/Context.ma".
inductive Weight: Nat \to Context \to Proof \to Nat \to Prop \def
include "datatypes_defs/Sequent.ma".
+theorem linj_inj: \forall a,c. linj a = linj c \to a = c.
+ intros; whd in H:(? ? % %); destruct; autobatch.
+qed.
+
theorem rinj_inj: \forall b,d. rinj b = rinj d \to b = d.
+ intros; whd in H:(? ? % %); destruct; autobatch.
+qed.
generalize in match Hn; generalize in match Hn; clear Hn;
unfold segment_enum;
generalize in match bound in ⊢ (% → ? → ? ? (? ? ? (? ? ? ? %)) ?);
- intros 1 (m); elim m (Hm Hn p IH Hm Hn); [ destruct Hm ]
+ intros 1 (m); elim m (Hm Hn p IH Hm Hn); [ simplify in Hm; destruct Hm ]
simplify; cases (eqP bool_eqType (ltb p bound) true); simplify;
[1:unfold segment in ⊢ (? ? match ? % ? ? with [_ ⇒ ?|_ ⇒ ?] ?);
unfold nat_eqType in ⊢ (? ? match % with [_ ⇒ ?|_ ⇒ ?] ?);
lemma uniqP : ∀d:eqType.∀l:list d.
reflect (∀x:d.mem d x l = true → count d (cmp d x) l = (S O)) (uniq d l).
-intros (d l); apply prove_reflect; elim l; [1: destruct H1 | 3: destruct H]
+intros (d l); apply prove_reflect; elim l; [1: simplify in H1; destruct H1 | 3: simplify in H; destruct H]
[1: generalize in match H2; simplify in H2;
lapply (b2pT ? ? (orbP ? ?) H2) as H3; clear H2;
cases H3; clear H3; intros;
unfold Not; intros (A); lapply (A t) as A';
[1: simplify in A'; rewrite > cmp_refl in A'; simplify in A';
destruct A'; rewrite < count_O_mem in H1;
- rewrite > Hcut in H1; destruct H1;
+ rewrite > Hcut in H1; simplify in H1; destruct H1;
|2: simplify; rewrite > cmp_refl; reflexivity;]
|2: intros (Ul1); lapply (H Ul1); unfold Not; intros (A); apply Hletin;
intros (r Mrl1); lapply (A r);
generalize in match Hletin1; simplify; apply (cmpP d r t);
simplify; intros (E Hc); [2: assumption]
destruct Hc; rewrite < count_O_mem in Mrl1;
- rewrite > Hcut in Mrl1; destruct Mrl1;]]
+ rewrite > Hcut in Mrl1; simplify in Mrl1; destruct Mrl1;]]
qed.
lemma mem_finType : ∀d:finType.∀x:d. mem d x (enum d) = true.
intros (d p x); cases x (t Ht); clear x;
generalize in match (mem_finType d t);
generalize in match (uniq_fintype_enum d);
-elim (enum d); [destruct H1] simplify;
+elim (enum d); [simplify in H1; destruct H1] simplify;
cases (in_sub_eq d p t1); simplify;
[1:generalize in match H3; clear H3; cases s (r Hr); clear s;
simplify; intros (Ert1); generalize in match Hr; clear Hr;
P l1 l2.
intros (T1 T2 l1 l2 P Hl Pnil Pcons);
generalize in match Hl; clear Hl; generalize in match l2; clear l2;
-elim l1 1 (l2 x1); [ cases l2; intros (Hl); [assumption| destruct Hl]]
+elim l1 1 (l2 x1); [ cases l2; intros (Hl); [assumption| simplify in Hl; destruct Hl]]
intros 3 (tl1 IH l2); cases l2; [1: simplify; intros 1 (Hl); destruct Hl]
-intros 1 (Hl); apply Pcons; apply IH; destruct Hl; assumption;
+intros 1 (Hl); apply Pcons; apply IH; simplify in Hl; destruct Hl; assumption;
qed.
lemma eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l.
lcmp ? l1 l2 = true → length ? l1 = length ? l2.
intros 2 (d l1); elim l1 1 (l2 x1);
[1: cases l2; simplify; intros; [reflexivity|destruct H]
-|2: intros 3 (tl1 IH l2); cases (l2); intros; [1:destruct H]
+|2: intros 3 (tl1 IH l2); cases (l2); intros; [1:simplify in H; destruct H]
simplify; (* XXX la apply non fa simplify? *)
apply congr_S; apply (IH l);
(* XXX qualcosa di enorme e' rotto! la regola di convertibilita?! *)
rewrite > (IH H2); rewrite > (b2pT ? ? (eqP d ? ?) H1); reflexivity
| generalize in match H; clear H; generalize in match l2; clear l2;
elim l1 1 (l1 x1);
- [ cases l1; [intros; destruct H | unfold Not; intros; destruct H1;]
+ [ cases l1; simplify; [intros; destruct H | unfold Not; intros; destruct H1;]
| intros 3 (tl1 IH l2); cases l2;
[ unfold Not; intros; destruct H1;
| simplify; intros;
|4,5:
simplify in H2;
destruct H2
- | simplify in H2;
+ | simplify in H1;
destruct H1
]
|4,5: