]> matita.cs.unibo.it Git - helm.git/commitdiff
- lambda: - normalization theorem completed!
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 10 Dec 2012 20:54:24 +0000 (20:54 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 10 Dec 2012 20:54:24 +0000 (20:54 +0000)
          - binary trees on pointer sequences set up to serve as labels for "st" computations
- lstar: right-oriened eliminator added

matita/matita/contribs/lambda/labelled_hap_computation.ma
matita/matita/contribs/lambda/pointer_tree.ma [new file with mode: 0644]
matita/matita/contribs/lambda/policy.txt
matita/matita/contribs/lambda/st_computation.ma
matita/matita/lib/basics/lists/lstar.ma

index 1fb6fcd180fec4084be345769ac517029f38c0bb..6309d6b5394022ec0091c60f97a144445e2cca3a 100644 (file)
@@ -70,6 +70,12 @@ theorem lhap_trans: ltransitive … lhap.
 /2 width=3 by lstar_ltransitive/
 qed-.
 
+lemma lhap_step_dx: ∀s,M1,M. M1 ⓗ⇀*[s] M →
+                    ∀p,M2. M ⓗ⇀[p] M2 → M1 ⓗ⇀*[s@p::◊] M2.
+#s #M1 #M #HM1 #p #M2 #HM2
+@(lhap_trans … HM1) /2 width=1/
+qed-.
+
 lemma head_lsreds_lhap: ∀s,M1,M2. M1 ⇀*[s] M2 → is_head s → M1 ⓗ⇀*[s] M2.
 #s #M1 #M2 #H @(lstar_ind_l ????????? H) -s -M1 //
 #a #s #M1 #M #HM1 #_ #IHM2 * /3 width=3/
diff --git a/matita/matita/contribs/lambda/pointer_tree.ma b/matita/matita/contribs/lambda/pointer_tree.ma
new file mode 100644 (file)
index 0000000..89d2122
--- /dev/null
@@ -0,0 +1,31 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "pointer_sequence.ma".
+
+(* POINTER TREE *************************************************************)
+
+(* Policy: pointer tree metavariables: P, Q *)
+(* Note: this is a binary tree on pointer sequences *)
+inductive ptree: Type[0] ≝
+| tnil : ptree
+| tcons: pseq → ptree → ptree → ptree
+.
+
+let rec size (P:ptree) on P ≝ match P with
+[ tnil          ⇒ 0
+| tcons s Q1 Q2 ⇒ size Q2 + size Q1 + |s|
+].
+
+interpretation "pointer tree size" 'card P = (size P).
index ef8c96fce518c20b368e1aa0e66604e09d0b2a6c..987d9f8b194b20631ea81e8f5d79d6ed18827407 100644 (file)
@@ -4,6 +4,7 @@ A, B, C, D: term
 H         : transient premise
 IH        : inductive premise
 M, N      : term
+P, Q      : pointer tree
 R         : arbitrary relation
 T, U      : arbitrary small type
 
index 97b5a7211658473f115e6dbd63487fc160335fb0..b536bb3a3708596bb4d9f8cd3b486475f5193bc3 100644 (file)
@@ -32,15 +32,96 @@ notation "hvbox( M ⓢ⥤* break term 46 N )"
    non associative with precedence 45
    for @{ 'Std $M $N }.
 
-axiom st_refl: reflexive … st.
+lemma st_inv_lref: ∀M,N. M ⓢ⥤* N → ∀j. #j = N →
+                   ∃s. M ⓗ⇀*[s] #j.
+#M #N * -M -N
+[ /2 width=2/
+| #s #M #A1 #A2 #_ #_ #j #H destruct
+| #s #M #B1 #B2 #A1 #A2 #_ #_ #_ #j #H destruct
+]
+qed-.
+
+lemma st_inv_abst: ∀M,N. M ⓢ⥤* N → ∀C2. 𝛌.C2 = N →
+                   ∃∃s,C1. M ⓗ⇀*[s] 𝛌.C1 & C1 ⓢ⥤* C2.
+#M #N * -M -N
+[ #s #M #i #_ #C2 #H destruct
+| #s #M #A1 #A2 #HM #A12 #C2 #H destruct /2 width=4/
+| #s #M #B1 #B2 #A1 #A2 #_ #_ #_ #C2 #H destruct
+]
+qed-.
+
+lemma st_inv_appl: ∀M,N. M ⓢ⥤* N → ∀D2,C2. @D2.C2 = N →
+                   ∃∃s,D1,C1. M ⓗ⇀*[s] @D1.C1 & D1 ⓢ⥤* D2 & C1 ⓢ⥤* C2.
+#M #N * -M -N
+[ #s #M #i #_ #D2 #C2 #H destruct
+| #s #M #A1 #A2 #_ #_ #D2 #C2 #H destruct 
+| #s #M #B1 #B2 #A1 #A2 #HM #HB12 #HA12 #D2 #C2 #H destruct /2 width=6/
+]
+qed-.
+
+lemma st_refl: reflexive … st.
+#M elim M -M /2 width=2/ /2 width=4/ /2 width=6/
+qed.
+
+lemma st_step_sn: ∀N1,N2. N1 ⓢ⥤* N2 → ∀s,M. M ⓗ⇀*[s] N1 → M ⓢ⥤* N2.
+#N1 #N2 #H elim H -N1 -N2
+[ #r #N #i #HN #s #M #HMN
+  lapply (lhap_trans … HMN … HN) -N /2 width=2/
+| #r #N #C1 #C2 #HN #_ #IHC12 #s #M #HMN
+  lapply (lhap_trans … HMN … HN) -N /3 width=5/
+| #r #N #D1 #D2 #C1 #C2 #HN #_ #_ #IHD12 #IHC12 #s #M #HMN
+  lapply (lhap_trans … HMN … HN) -N /3 width=7/
+]
+qed-.
 
-axiom st_step_sn: ∀N1,N2. N1 ⓢ⥤* N2 → ∀s,M. M ⓗ⇀*[s] N1 → M ⓢ⥤* N2.
+lemma st_step_rc: ∀s,M1,M2. M1 ⓗ⇀*[s] M2 → M1 ⓢ⥤* M2.
+/3 width=4 by st_step_sn/
+qed.
 
-axiom st_lift: liftable st.
+lemma st_lift: liftable st.
+#h #M1 #M2 #H elim H -M1 -M2
+[ /3 width=2/
+| #s #M #A1 #A2 #HM #_ #IHA12 #d
+  @st_abst [3: @(lhap_lift … HM) |1,2: skip ] -M // (**) (* auto fails here *)
+| #s #M #B1 #B2 #A1 #A2 #HM #_ #_ #IHB12 #IHA12 #d
+  @st_appl [4: @(lhap_lift … HM) |1,2,3: skip ] -M // (**) (* auto fails here *)
+]
+qed.
 
-axiom st_inv_lift: deliftable_sn st.
+lemma st_inv_lift: deliftable_sn st.
+#h #N1 #N2 #H elim H -N1 -N2
+[ #s #N1 #i #HN1 #d #M1 #HMN1
+  elim (lhap_inv_lift … HN1 … HMN1) -N1 /3 width=3/
+| #s #N1 #C1 #C2 #HN1 #_ #IHC12 #d #M1 #HMN1
+  elim (lhap_inv_lift … HN1 … HMN1) -N1 #M2 #HM12 #HM2
+  elim (lift_inv_abst … HM2) -HM2 #A1 #HAC1 #HM2 destruct
+  elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #HAC2 destruct (**) (* simplify line *)
+  @(ex2_intro … (𝛌.A2)) // /2 width=4/
+| #s #N1 #D1 #D2 #C1 #C2 #HN1 #_ #_ #IHD12 #IHC12 #d #M1 #HMN1
+  elim (lhap_inv_lift … HN1 … HMN1) -N1 #M2 #HM12 #HM2
+  elim (lift_inv_appl … HM2) -HM2 #B1 #A1 #HBD1 #HAC1 #HM2 destruct
+  elim (IHD12 ???) -IHD12 [4: // |2,3: skip ] #B2 #HB12 #HBD2 destruct (**) (* simplify line *)
+  elim (IHC12 ???) -IHC12 [4: // |2,3: skip ] #A2 #HA12 #HAC2 destruct (**) (* simplify line *)
+  @(ex2_intro … (@B2.A2)) // /2 width=6/
+]
+qed-.
 
-axiom st_dsubst: dsubstable st.
+lemma st_dsubst: dsubstable st.
+#N1 #N2 #HN12 #M1 #M2 #H elim H -M1 -M2
+[ #s #M #i #HM #d elim (lt_or_eq_or_gt i d) #Hid
+  [ lapply (lhap_dsubst … N1 … HM d) -HM
+    >(dsubst_vref_lt … Hid) >(dsubst_vref_lt … Hid) /2 width=2/
+  | destruct >dsubst_vref_eq
+    @(st_step_sn (↑[0,i]N1) … s) /2 width=1/
+  | lapply (lhap_dsubst … N1 … HM d) -HM
+    >(dsubst_vref_gt … Hid) >(dsubst_vref_gt … Hid) /2 width=2/
+  ]
+| #s #M #A1 #A2 #HM #_ #IHA12 #d
+  lapply (lhap_dsubst … N1 … HM d) -HM /2 width=4/ (**) (* auto needs some help here *)
+| #s #M #B1 #B2 #A1 #A2 #HM #_ #_ #IHB12 #IHA12 #d
+  lapply (lhap_dsubst … N1 … HM d) -HM /2 width=6/ (**) (* auto needs some help here *)
+]
+qed.
 
 lemma st_inv_lsreds_is_le: ∀M,N. M ⓢ⥤* N →
                            ∃∃r. M ⇀*[r] N & is_le r.
@@ -62,3 +143,37 @@ lemma st_inv_lsreds_is_le: ∀M,N. M ⓢ⥤* N →
   @(ex2_intro … HM) -M -B2 -A2 >associative_append /3 width=1/
 ]
 qed-.
+
+lemma st_step_dx: ∀p,M,M2. M ⇀[p] M2 → ∀M1. M1 ⓢ⥤* M → M1 ⓢ⥤* M2.
+#p #M #M2 #H elim H -p -M -M2
+[ #B #A #M1 #H
+  elim (st_inv_appl … H ???) -H [4: // |2,3: skip ] #s #B1 #M #HM1 #HB1 #H (**) (* simplify line *)
+  elim (st_inv_abst … H ??) -H [3: // |2: skip ] #r #A1 #HM #HA1 (**) (* simplify line *)
+  @(st_step_sn … ([⬐B1]A1) … (s@(dx:::r)@(◊::◊))) /2 width=1/ -B -A
+  @(lhap_trans … HM1) -M1
+  @(lhap_step_dx … (@B1.𝛌.A1)) // -s /2 width=1/
+| #p #A #A2 #_ #IHA2 #M1 #H
+  elim (st_inv_abst … H ??) -H [3: // |2: skip ] /3 width=4/ (**) (* simplify line *)
+| #p #B #B2 #A #_ #IHB2 #M1 #H
+  elim (st_inv_appl … H ???) -H [4: // |2,3: skip ] /3 width=6/ (**) (* simplify line *)
+| #p #B #A #A2 #_ #IHA2 #M1 #H
+  elim (st_inv_appl … H ???) -H [4: // |2,3: skip ] /3 width=6/ (**) (* simplify line *)
+]
+qed-.
+
+lemma st_lsreds: ∀s,M1,M2. M1 ⇀*[s] M2 → M1 ⓢ⥤* M2.
+#s #M1 #M2 #H @(lstar_ind_r ????????? H) -s -M2 // /2 width=4 by st_step_dx/
+qed.
+
+theorem st_trans: transitive … st.
+#M1 #M #M2 #HM1 #HM2
+elim (st_inv_lsreds_is_le … HM1) -HM1 #s1 #HM1 #_
+elim (st_inv_lsreds_is_le … HM2) -HM2 #s2 #HM2 #_
+lapply (lsreds_trans … HM1 … HM2) -M /2 width=2/
+qed-.
+
+theorem lsreds_standard: ∀s,M,N. M ⇀*[s] N →
+                         ∃∃r. M ⇀*[r] N & is_le r.
+#s #M #N #H
+@st_inv_lsreds_is_le /2 width=2/
+qed-.
index fa0cdb5aaf4f570b8b4fc02d3f76038dab7300ba..780f7759dd31f2d8d43217e391abad68d946250e 100644 (file)
@@ -78,3 +78,48 @@ qed-.
 theorem lstar_ltransitive: ∀A,B,R. ltransitive … (lstar A B R).
 #A #B #R #l1 #b1 #b #H @(lstar_ind_l ????????? H) -l1 -b1 normalize // /3 width=3/
 qed-.
+
+lemma lstar_app: ∀A,B,R,l,b1,b. lstar A B R l b1 b → ∀a,b2. R a b b2 →
+                 lstar A B R (l@[a]) b1 b2.
+#A #B #R #l #b1 #b #H @(lstar_ind_l ????????? H) -l -b1 /2 width=1/
+normalize /3 width=3/
+qed.
+
+inductive lstar_r (A:Type[0]) (B:Type[0]) (R: A→relation B): list A → relation B ≝
+| lstar_r_nil: ∀b. lstar_r A B R ([]) b b
+| lstar_r_app: ∀l,b1,b. lstar_r A B R l b1 b → ∀a,b2. R a b b2 →
+               lstar_r A B R (l@[a]) b1 b2
+.
+
+lemma lstar_r_cons: ∀A,B,R,l,b,b2. lstar_r A B R l b b2 → ∀a,b1. R a b1 b →
+                    lstar_r A B R (a::l) b1 b2.
+#A #B #R #l #b #b2 #H elim H -l -b2 /2 width=3/
+#l #b1 #b #_ #a #b2 #Hb2 #IHb1 #a0 #b0 #Hb01
+@(lstar_r_app … (a0::l) … Hb2) -b2 /2 width=1/
+qed.
+
+lemma lstar_lstar_r: ∀A,B,R,l,b1,b2. lstar A B R l b1 b2 → lstar_r A B R l b1 b2.
+#A #B #R #l #b1 #b2 #H @(lstar_ind_l ????????? H) -l -b1 // /2 width=3/
+qed.
+
+lemma lstar_r_inv_lstar: ∀A,B,R,l,b1,b2. lstar_r A B R l b1 b2 → lstar A B R l b1 b2.
+#A #B #R #l #b1 #b2 #H elim H -l -b1 -b2 // /2 width=3/
+qed-.
+
+fact lstar_ind_r_aux: ∀A,B,R,b1. ∀P:relation2 (list A) B.
+                      P ([]) b1 →
+                      (∀a,l,b,b2. lstar … R l b1 b → R a b b2 → P l b → P (l@[a]) b2) →
+                      ∀l,b,b2. lstar … R l b b2 → b = b1 → P l b2.
+#A #B #R #b1 #P #H1 #H2 #l #b #b2 #H elim (lstar_lstar_r ?????? H) -l -b -b2
+[ #b #H destruct //
+| #l #b #b0 #Hb0 #a #b2 #Hb02 #IH #H destruct /3 width=4 by lstar_r_inv_lstar/
+]
+qed-.
+
+lemma lstar_ind_r: ∀A,B,R,b1. ∀P:relation2 (list A) B.
+                   P ([]) b1 →
+                   (∀a,l,b,b2. lstar … R l b1 b → R a b b2 → P l b → P (l@[a]) b2) →
+                   ∀l,b2. lstar … R l b1 b2 → P l b2.
+#A #B #R #b1 #P #H1 #H2 #l #b2 #Hb12
+@(lstar_ind_r_aux … H1 H2 … Hb12) //
+qed-.