]> matita.cs.unibo.it Git - helm.git/commitdiff
- ground_2: support for lifts_div4
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 18 May 2016 18:44:46 +0000 (18:44 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 18 May 2016 18:44:46 +0000 (18:44 +0000)
- basic_2: cpg_drops competed

14 files changed:
matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_drops.etc [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_length.etc [new file with mode: 0644]
matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/lifts.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_lifts.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_length.ma [deleted file]
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_lsubr.ma
matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_simple.ma
matita/matita/contribs/lambdadelta/basic_2/static/aaa_drops.ma
matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_after.ma
matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_after.ma
matita/matita/contribs/lambdadelta/ground_2/relocation/rtmap_at.ma

diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_drops.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_drops.etc
new file mode 100644 (file)
index 0000000..c0447ff
--- /dev/null
@@ -0,0 +1,16 @@
+lemma cpg_delift: ∀c,h,I,G,K,V,T1,L,i. ⬇*[i] L ≡ (K.ⓑ{I}V) →
+                  ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, 𝟘𝟘] T2 & ⬆*[↑1] T ≡ T2.
+#h #c #I #G #K #V #T1 elim T1 -T1
+[ * #i #L #l /2 width=4 by cpg_atom, lift_sort, lift_gref, ex2_2_intro/
+  elim (lt_or_eq_or_gt i l) #Hil [1,3: /4 width=4 by cpg_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ]
+  destruct
+  elim (lift_total V 0 (i+1)) #W #HVW
+  elim (lift_split … HVW i i) /3 width=7 by cpg_delta, ex2_2_intro/
+| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK
+  elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
+  [ elim (IHU1 (L. ⓑ{I} W1) (l+1)) -IHU1 /3 width=9 by cpg_bind, drop_drop, lift_bind, ex2_2_intro/
+  | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpg_flat, lift_flat, ex2_2_intro/
+  ]
+]
+qed-.
+*)
diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_length.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpg/cpg_length.etc
new file mode 100644 (file)
index 0000000..142f0fd
--- /dev/null
@@ -0,0 +1,26 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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 "basic_2/relocation/drops.ma".
+include "basic_2/rt_transition/cpg.ma".
+
+(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
+
+(* Properties with length for local environments ****************************)
+
+lemma cpg_inv_lref1_ge: ∀h,c,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, c] T2 → |L| ≤ i → T2 = #i.
+#h #c #G #L #T2 #i #H elim (cpg_inv_lref1 … H) -H // *
+#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
+#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+qed-.
index 06a2eb84e10399a76eb88695661943561215bba6..0b3279de95214bf6caea3b889a5e0fd858bd758e 100644 (file)
@@ -31,7 +31,7 @@ theorem drops_conf: ∀b1,f1,L1,L. ⬇*[b1, f1] L1 ≡ L →
   #g #Hg #H destruct /3 width=3 by drops_inv_drop1/
 | #f1 #I #K1 #K #V1 #V #_ #HV1 #IH #b2 #f #L2 #HL2 #f2 #Hf elim (after_inv_pxx … Hf) -Hf [1,3: * |*:// ]
   #g2 #g #Hf #H1 #H2 destruct
-  [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_div/
+  [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_div3/
   | /4 width=3 by drops_inv_drop1, drops_drop/
   ]
 ]
index ceaca58a944a6c4884fd476b000b8f2315bebcd2..495d5e45fa9cbddac708ae4aa0f836db860e484d 100644 (file)
@@ -208,6 +208,8 @@ lemma lifts_inv_flat2: ∀f:rtmap. ∀I,V2,T2,X. ⬆*[f] X ≡ ⓕ{I}V2.T2 →
                                 X = ⓕ{I}V1.T1.
 /2 width=3 by lifts_inv_flat2_aux/ qed-.
 
+(* Advanced inversion lemmas ************************************************)
+
 (* Basic_2A1: includes: lift_inv_pair_xy_x *)
 lemma lifts_inv_pair_xy_x: ∀f,I,V,T. ⬆*[f] ②{I}V.T ≡ V → ⊥.
 #f #J #V elim V -V
@@ -241,6 +243,10 @@ lemma lifts_inv_pair_xy_y: ∀I,T,V,f. ⬆*[f] ②{I}V.T ≡ T → ⊥.
 ]
 qed-.
 
+lemma lifts_inv_lref1_uni: ∀l,Y,i. ⬆*[l] #i ≡ Y → Y = #(l+i).
+#l #Y #i1 #H elim (lifts_inv_lref1 … H) -H /4 width=4 by at_mono, eq_f/
+qed-.
+
 (* Basic forward lemmas *****************************************************)
 
 (* Basic_2A1: includes: lift_inv_O2 *)
@@ -371,6 +377,11 @@ lemma is_lifts_dec: ∀T2,f. Decidable (∃T1. ⬆*[f] T1 ≡ T2).
 ]
 qed-.
 
+(* Properties with uniform relocation ***************************************)
+
+lemma lifts_uni: ∀n1,n2,T,U. ⬆*[𝐔❴n1❵∘𝐔❴n2❵] T ≡ U → ⬆*[n1+n2] T ≡ U.
+/3 width=4 by lifts_eq_repl_back, after_inv_total/ qed.
+
 (* Basic_2A1: removed theorems 14:
               lifts_inv_nil lifts_inv_cons
               lift_inv_Y1 lift_inv_Y2 lift_inv_lref_Y1 lift_inv_lref_Y2 lift_lref_Y lift_Y1
index 79c2000872024cb32a4282b396a39fe88055a84d..8477714f3d2c9ff7e30c3b8e2b319523a6bbdb88 100644 (file)
@@ -20,17 +20,48 @@ include "basic_2/relocation/lifts.ma".
 
 (* Basic_1: includes: lift_gen_lift *)
 (* Basic_2A1: includes: lift_div_le lift_div_be *)
-theorem lifts_div: ∀f2,T,T2. ⬆*[f2] T2 ≡ T → ∀T1,f. ⬆*[f] T1 ≡ T →
-                   ∀f1. f2 ⊚ f1 ≡ f → ⬆*[f1] T1 ≡ T2.
+theorem lift_div4: ∀f2,Tf,T. ⬆*[f2] Tf ≡ T → ∀g2,Tg. ⬆*[g2] Tg ≡ T →
+                   ∀f1,g1. H_at_div f2 g2 f1 g1 →
+                   ∃∃T0. ⬆*[f1] T0 ≡ Tf & ⬆*[g1] T0 ≡ Tg.
+#f2 #Tf #T #H elim H -f2 -Tf -T
+[ #f2 #s #g2 #Tg #H #f1 #g1 #_
+  lapply (lifts_inv_sort2 … H) -H #H destruct
+  /2 width=3 by ex2_intro/
+| #f2 #jf #j #Hf2 #g2 #Tg #H #f1 #g1 #H0
+  elim (lifts_inv_lref2 … H) -H #jg #Hg2 #H destruct
+  elim (H0 … Hf2 Hg2) -H0 -j /3 width=3 by lifts_lref, ex2_intro/
+| #f2 #l #g2 #Tg #H #f1 #g1 #_
+  lapply (lifts_inv_gref2 … H) -H #H destruct
+  /2 width=3 by ex2_intro/
+| #f2 #p #I #Vf #V #Tf #T #_ #_ #IHV #IHT #g2 #X #H #f1 #g1 #H0
+  elim (lifts_inv_bind2 … H) -H #Vg #Tg #HVg #HTg #H destruct
+  elim (IHV … HVg … H0) -IHV -HVg
+  elim (IHT … HTg) -IHT -HTg [ |*: /2 width=8 by at_div_pp/ ]
+  /3 width=5 by lifts_bind, ex2_intro/
+| #f2 #I #Vf #V #Tf #T #_ #_ #IHV #IHT #g2 #X #H #f1 #g1 #H0
+  elim (lifts_inv_flat2 … H) -H #Vg #Tg #HVg #HTg #H destruct
+  elim (IHV … HVg … H0) -IHV -HVg
+  elim (IHT … HTg … H0) -IHT -HTg -H0
+  /3 width=5 by lifts_flat, ex2_intro/
+]
+qed-.
+
+lemma lifts_div4_one: ∀f,Tf,T. ⬆*[↑f] Tf ≡ T →
+                      ∀T1. ⬆*[1] T1 ≡ T →
+                      ∃∃T0. ⬆*[1] T0 ≡ Tf & ⬆*[f] T0 ≡ T1.
+/4 width=6 by lift_div4, at_div_id_dx, at_div_pn/ qed-.
+
+theorem lifts_div3: ∀f2,T,T2. ⬆*[f2] T2 ≡ T → ∀f,T1. ⬆*[f] T1 ≡ T →
+                    ∀f1. f2 ⊚ f1 ≡ f → ⬆*[f1] T1 ≡ T2.
 #f2 #T #T2 #H elim H -f2 -T -T2
-[ #f2 #s #T1 #f #H >(lifts_inv_sort2 … H) -T1 //
-| #f2 #i2 #i #Hi2 #T1 #f #H #f1 #Ht21 elim (lifts_inv_lref2 … H) -H
+[ #f2 #s #f #T1 #H >(lifts_inv_sort2 … H) -T1 //
+| #f2 #i2 #i #Hi2 #f #T1 #H #f1 #Ht21 elim (lifts_inv_lref2 … H) -H
   #i1 #Hi1 #H destruct /3 width=6 by lifts_lref, after_fwd_at1/
-| #f2 #l #T1 #f #H >(lifts_inv_gref2 … H) -T1 //
-| #f2 #p #I #W2 #W #U2 #U #_ #_ #IHW #IHU #T1 #f #H
+| #f2 #l #f #T1 #H >(lifts_inv_gref2 … H) -T1 //
+| #f2 #p #I #W2 #W #U2 #U #_ #_ #IHW #IHU #f #T1 #H
   elim (lifts_inv_bind2 … H) -H #W1 #U1 #HW1 #HU1 #H destruct
   /4 width=3 by lifts_bind, after_O2/
-| #f2 #I #W2 #W #U2 #U #_ #_ #IHW #IHU #T1 #f #H
+| #f2 #I #W2 #W #U2 #U #_ #_ #IHW #IHU #f #T1 #H
   elim (lifts_inv_flat2 … H) -H #W1 #U1 #HW1 #HU1 #H destruct
   /3 width=3 by lifts_flat/
 ]
@@ -39,34 +70,34 @@ qed-.
 (* Basic_1: was: lift1_lift1 (left to right) *)
 (* Basic_1: includes: lift_free (left to right) lift_d lift1_xhg (right to left) lift1_free (right to left) *)
 (* Basic_2A1: includes: lift_trans_be lift_trans_le lift_trans_ge lifts_lift_trans_le lifts_lift_trans *)
-theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≡ T → ∀T2,f2. ⬆*[f2] T ≡ T2 →
+theorem lifts_trans: ∀f1,T1,T. ⬆*[f1] T1 ≡ T → ∀f2,T2. ⬆*[f2] T ≡ T2 →
                      ∀f. f2 ⊚ f1 ≡ f → ⬆*[f] T1 ≡ T2.
 #f1 #T1 #T #H elim H -f1 -T1 -T
-[ #f1 #s #T2 #f2 #H >(lifts_inv_sort1 … H) -T2 //
-| #f1 #i1 #i #Hi1 #T2 #f2 #H #f #Ht21 elim (lifts_inv_lref1 … H) -H
+[ #f1 #s #f2 #T2 #H >(lifts_inv_sort1 … H) -T2 //
+| #f1 #i1 #i #Hi1 #f2 #T2 #H #f #Ht21 elim (lifts_inv_lref1 … H) -H
   #i2 #Hi2 #H destruct /3 width=6 by lifts_lref, after_fwd_at/
-| #f1 #l #T2 #f2 #H >(lifts_inv_gref1 … H) -T2 //
-| #f1 #p #I #W1 #W #U1 #U #_ #_ #IHW #IHU #T2 #f2 #H
+| #f1 #l #f2 #T2 #H >(lifts_inv_gref1 … H) -T2 //
+| #f1 #p #I #W1 #W #U1 #U #_ #_ #IHW #IHU #f2 #T2 #H
   elim (lifts_inv_bind1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
   /4 width=3 by lifts_bind, after_O2/
-| #f1 #I #W1 #W #U1 #U #_ #_ #IHW #IHU #T2 #f2 #H
+| #f1 #I #W1 #W #U1 #U #_ #_ #IHW #IHU #f2 #T2 #H
   elim (lifts_inv_flat1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
   /3 width=3 by lifts_flat/
 ]
 qed-.
 
 (* Basic_2A1: includes: lift_conf_O1 lift_conf_be *)
-theorem lifts_conf: ∀f1,T,T1. ⬆*[f1] T ≡ T1 → ∀T2,f. ⬆*[f] T ≡ T2 →
+theorem lifts_conf: ∀f1,T,T1. ⬆*[f1] T ≡ T1 → ∀f,T2. ⬆*[f] T ≡ T2 →
                     ∀f2. f2 ⊚ f1 ≡ f → ⬆*[f2] T1 ≡ T2.
 #f1 #T #T1 #H elim H -f1 -T -T1
-[ #f1 #s #T2 #f #H >(lifts_inv_sort1 … H) -T2 //
-| #f1 #i #i1 #Hi1 #T2 #f #H #f2 #Ht21 elim (lifts_inv_lref1 … H) -H
+[ #f1 #s #f #T2 #H >(lifts_inv_sort1 … H) -T2 //
+| #f1 #i #i1 #Hi1 #f #T2 #H #f2 #Ht21 elim (lifts_inv_lref1 … H) -H
   #i2 #Hi2 #H destruct /3 width=6 by lifts_lref, after_fwd_at2/
-| #f1 #l #T2 #f #H >(lifts_inv_gref1 … H) -T2 //
-| #f1 #p #I #W #W1 #U #U1 #_ #_ #IHW #IHU #T2 #f #H
+| #f1 #l #f #T2 #H >(lifts_inv_gref1 … H) -T2 //
+| #f1 #p #I #W #W1 #U #U1 #_ #_ #IHW #IHU #f #T2 #H
   elim (lifts_inv_bind1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
   /4 width=3 by lifts_bind, after_O2/
-| #f1 #I #W #W1 #U #U1 #_ #_ #IHW #IHU #T2 #f #H
+| #f1 #I #W #W1 #U #U1 #_ #_ #IHW #IHU #f #T2 #H
   elim (lifts_inv_flat1 … H) -H #W2 #U2 #HW2 #HU2 #H destruct
   /3 width=3 by lifts_flat/
 ]
@@ -77,7 +108,7 @@ qed-.
 (* Basic_2A1: includes: lift_inj *)
 lemma lifts_inj: ∀f,T1,U. ⬆*[f] T1 ≡ U → ∀T2. ⬆*[f] T2 ≡ U → T1 = T2.
 #f #T1 #U #H1 #T2 #H2 lapply (isid_after_dx 𝐈𝐝  … f)
-/3 width=6 by lifts_div, lifts_fwd_isid/
+/3 width=6 by lifts_div3, lifts_fwd_isid/
 qed-.
 
 (* Basic_2A1: includes: lift_mono *)
index 12c45c9feea984308225a25786098c9c342e9e92..22c671d90e4aeb3afeb82a2ff73b5bcd34905a8a 100644 (file)
@@ -30,7 +30,7 @@ inductive cpg (h): rtc → relation4 genv lenv term term ≝
              ⬆*[1] V2 ≡ W2 → cpg h (↓c) G (L.ⓓV1) (#0) W2
 | cpg_ell  : ∀c,G,L,V1,V2,W2. cpg h c G L V1 V2 →
              ⬆*[1] V2 ≡ W2 → cpg h ((↓c)+𝟘𝟙) G (L.ⓛV1) (#0) W2
-| cpt_lref : ∀c,I,G,L,V,T,U,i. cpg h c G L (#i) T → 
+| cpg_lref : ∀c,I,G,L,V,T,U,i. cpg h c G L (#i) T → 
              ⬆*[1] T ≡ U → cpg h c G (L.ⓑ{I}V) (#⫯i) U
 | cpg_bind : ∀cV,cT,p,I,G,L,V1,V2,T1,T2.
              cpg h cV G L V1 V2 → cpg h cT G (L.ⓑ{I}V1) T1 T2 →
@@ -70,8 +70,8 @@ qed.
 (* Basic inversion lemmas ***************************************************)
 
 fact cpg_inv_atom1_aux: ∀c,h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[c, h] T2 → ∀J. T1 = ⓪{J} →
-                        ∨∨ T2 = ⓪{J}
-                         | ∃∃s. J = Sort s & T2 = ⋆(next h s)
+                        ∨∨ T2 = ⓪{J} ∧ c = 𝟘𝟘 
+                         | ∃∃s. J = Sort s & T2 = ⋆(next h s) & c = 𝟘𝟙
                          | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
                                          L = K.ⓓV1 & J = LRef 0 & c = ↓cV
                          | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
@@ -79,8 +79,8 @@ fact cpg_inv_atom1_aux: ∀c,h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[c, h] T2 → ∀
                          | ∃∃I,K,V,T,i. ⦃G, K⦄ ⊢ #i ➡[c, h] T & ⬆*[1] T ≡ T2 &
                                         L = K.ⓑ{I}V & J = LRef (⫯i).
 #c #h #G #L #T1 #T2 * -c -G -L -T1 -T2
-[ #I #G #L #J #H destruct /2 width=1 by or5_intro0/
-| #G #L #s #J #H destruct /3 width=3 by or5_intro1, ex2_intro/
+[ #I #G #L #J #H destruct /3 width=1 by or5_intro0, conj/
+| #G #L #s #J #H destruct /3 width=3 by or5_intro1, ex3_intro/
 | #c #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro2, ex5_4_intro/
 | #c #G #L #V1 #V2 #W2 #HV12 #VW2 #J #H destruct /3 width=8 by or5_intro3, ex5_4_intro/
 | #c #I #G #L #V #T #U #i #HT #HTU #J #H destruct /3 width=9 by or5_intro4, ex4_5_intro/
@@ -95,8 +95,8 @@ fact cpg_inv_atom1_aux: ∀c,h,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[c, h] T2 → ∀
 qed-.
 
 lemma cpg_inv_atom1: ∀c,h,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[c, h] T2 →
-                     ∨∨ T2 = ⓪{J}
-                      | ∃∃s. J = Sort s & T2 = ⋆(next h s)
+                     ∨∨ T2 = ⓪{J} ∧ c = 𝟘𝟘 
+                      | ∃∃s. J = Sort s & T2 = ⋆(next h s) & c = 𝟘𝟙
                       | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
                                       L = K.ⓓV1 & J = LRef 0 & c = ↓cV
                       | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
@@ -106,23 +106,23 @@ lemma cpg_inv_atom1: ∀c,h,J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[c, h] T2 →
 /2 width=3 by cpg_inv_atom1_aux/ qed-.
 
 lemma cpg_inv_sort1: ∀c,h,G,L,T2,s. ⦃G, L⦄ ⊢ ⋆s ➡[c, h] T2 →
-                     T2 = ⋆s ∨ T2 = ⋆(next h s).
+                     (T2 = ⋆s ∧ c = 𝟘𝟘) ∨ (T2 = ⋆(next h s) ∧ c = 𝟘𝟙).
 #c #h #G #L #T2 #s #H
-elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #s0 #H destruct /2 width=1 by or_intror/
+elim (cpg_inv_atom1 … H) -H * /3 width=1 by or_introl, conj/
+[ #s0 #H destruct /3 width=1 by or_intror, conj/
 |2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
 | #I #K #V1 #V2 #i #_ #_ #_ #H destruct
 ]
 qed-.
 
 lemma cpg_inv_zero1: ∀c,h,G,L,T2. ⦃G, L⦄ ⊢ #0 ➡[c, h] T2 →
-                     ∨∨ T2 = #0 
+                     ∨∨ (T2 = #0 ∧ c = 𝟘𝟘)
                       | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
                                       L = K.ⓓV1 & c = ↓cV
                       | ∃∃cV,K,V1,V2. ⦃G, K⦄ ⊢ V1 ➡[cV, h] V2 & ⬆*[1] V2 ≡ T2 &
                                       L = K.ⓛV1 & c = (↓cV)+𝟘𝟙.
 #c #h #G #L #T2 #H
-elim (cpg_inv_atom1 … H) -H /2 width=1 by or3_intro0/ *
+elim (cpg_inv_atom1 … H) -H * /3 width=1 by or3_intro0, conj/
 [ #s #H destruct
 |2,3: #cV #K #V1 #V2 #HV12 #HVT2 #H1 #_ #H2 destruct /3 width=8 by or3_intro1, or3_intro2, ex4_4_intro/
 | #I #K #V1 #V2 #i #_ #_ #_ #H destruct
@@ -130,19 +130,19 @@ elim (cpg_inv_atom1 … H) -H /2 width=1 by or3_intro0/ *
 qed-.
 
 lemma cpg_inv_lref1: ∀c,h,G,L,T2,i. ⦃G, L⦄ ⊢ #⫯i ➡[c, h] T2 →
-                     (T2 = #⫯i) ∨
+                     (T2 = #(⫯i) ∧ c = 𝟘𝟘) ∨
                      ∃∃I,K,V,T. ⦃G, K⦄ ⊢ #i ➡[c, h] T & ⬆*[1] T ≡ T2 & L = K.ⓑ{I}V.
 #c #h #G #L #T2 #i #H
-elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
+elim (cpg_inv_atom1 … H) -H * /3 width=1 by or_introl, conj/
 [ #s #H destruct
 |2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
 | #I #K #V1 #V2 #j #HV2 #HVT2 #H1 #H2 destruct /3 width=7 by ex3_4_intro, or_intror/
 ]
 qed-.
 
-lemma cpg_inv_gref1: ∀c,h,G,L,T2,l.  ⦃G, L⦄ ⊢ §l ➡[c, h] T2 → T2 = §l.
+lemma cpg_inv_gref1: ∀c,h,G,L,T2,l. ⦃G, L⦄ ⊢ §l ➡[c, h] T2 → T2 = §l ∧ c = 𝟘𝟘.
 #c #h #G #L #T2 #l #H
-elim (cpg_inv_atom1 … H) -H // *
+elim (cpg_inv_atom1 … H) -H * /2 width=1 by conj/
 [ #s #H destruct
 |2,3: #cV #K #V1 #V2 #_ #_ #_ #H destruct
 | #I #K #V1 #V2 #i #_ #_ #_ #H destruct
@@ -190,9 +190,9 @@ lemma cpg_inv_abbr1: ∀c,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{p}V1.T1 ➡[c, h]
 /3 width=8 by ex4_4_intro, ex4_2_intro, or_introl, or_intror/
 qed-.
 
-lemma cpg_inv_abst1: ∀c,h,p,G,L,V1,T1,U2.  ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[c, h] U2 →
+lemma cpg_inv_abst1: ∀c,h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡[c, h] U2 →
                      ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡[cT, h] T2 &
-                                    U2 = ⓛ{p} V2. T2 & c = (↓cV)+cT.
+                                    U2 = ⓛ{p}V2.T2 & c = (↓cV)+cT.
 #c #h #p #G #L #V1 #T1 #U2 #H elim (cpg_inv_bind1 … H) -H * 
 [ /3 width=8 by ex4_4_intro/
 | #c #T #_ #_ #_ #H destruct
index 6a235c4d9b18f20d43cb4cc2cba81726452ad790..4ccf66fdfa32ea49c65e90627b36a7a72d5b6578 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/drops.ma".
+include "basic_2/relocation/drops_drops.ma".
+include "basic_2/s_computation/fqup_weight.ma".
+include "basic_2/s_computation/fqup_drops.ma".
 include "basic_2/rt_transition/cpg.ma".
 
 (* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
 
+(* Advanced properties ******************************************************)
+
+lemma cpg_delta_drops: ∀c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡[c, h] V2 →
+                       ⬆*[⫯i] V2 ≡ T2 →  ⦃G, L⦄ ⊢ #i ➡[↓c, h] T2.
+#c #h #G #K #V #V2 #i elim i -i
+[ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_delta/
+| #i #IH #L0 #T0 #H0 #HV2 #HVT2
+  elim (drops_inv_succ … H0) -H0 #I #L #V0 #HLK #H destruct
+  elim (lifts_split_trans … HVT2 (𝐔❴⫯i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
+]
+qed.
+
+lemma cpg_ell_drops: ∀c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≡ K.ⓛV → ⦃G, K⦄ ⊢ V ➡[c, h] V2 →
+                     ⬆*[⫯i] V2 ≡ T2 →  ⦃G, L⦄ ⊢ #i ➡[(↓c)+𝟘𝟙, h] T2.
+#c #h #G #K #V #V2 #i elim i -i
+[ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_ell/
+| #i #IH #L0 #T0 #H0 #HV2 #HVT2
+  elim (drops_inv_succ … H0) -H0 #I #L #V0 #HLK #H destruct
+  elim (lifts_split_trans … HVT2 (𝐔❴⫯i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
+]
+qed.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma cpg_inv_lref1_drops: ∀c,h,G,i,L,T2. ⦃G, L⦄ ⊢ #i ➡[c, h] T2 →
+                           ∨∨ T2 = #i ∧ c = 𝟘𝟘
+                            | ∃∃cV,K,V,V2. ⬇*[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V ➡[cV, h] V2 &
+                                           ⬆*[⫯i] V2 ≡ T2 & c = ↓cV
+                            | ∃∃cV,K,V,V2. ⬇*[i] L ≡ K.ⓛV & ⦃G, K⦄ ⊢ V ➡[cV, h] V2 &
+                                           ⬆*[⫯i] V2 ≡ T2 & c = (↓cV) + 𝟘𝟙.
+#c #h #G #i elim i -i
+[ #L #T2 #H elim (cpg_inv_zero1 … H) -H * /3 width=1 by or3_intro0, conj/
+  /4 width=8 by drops_refl, ex4_4_intro, or3_intro2, or3_intro1/
+| #i #IH #L #T2 #H elim (cpg_inv_lref1 … H) -H * /3 width=1 by or3_intro0, conj/
+  #I #K #V #V2 #H #HVT2 #H0 destruct elim (IH … H) -IH -H
+  [ * #H1 #H2 destruct lapply (lifts_inv_lref1_uni … HVT2) -HVT2 #H destruct /3 width=1 by or3_intro0, conj/ ] *
+  #cV #L #W #W2 #HKL #HW2 #HWV2 #H destruct
+  lapply (lifts_trans … HWV2 … HVT2 ??) -V2
+  /4 width=8 by drops_drop, ex4_4_intro, or3_intro2, or3_intro1/
+]
+qed-.
+
 (* Properties with generic slicing for local environments *******************)
 
-lemma cpg_delift: ∀c,h,I,G,K,V,T1,L,i. ⬇*[i] L ≡ (K.ⓑ{I}V) →
-                  ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, 𝟘𝟘] T2 & ⬆*[↑1] T ≡ T2.
-#h #c #I #G #K #V #T1 elim T1 -T1
-[ * #i #L #l /2 width=4 by cpg_atom, lift_sort, lift_gref, ex2_2_intro/
-  elim (lt_or_eq_or_gt i l) #Hil [1,3: /4 width=4 by cpg_atom, lift_lref_ge_minus, lift_lref_lt, ylt_inj, yle_inj, ex2_2_intro/ ]
-  destruct
-  elim (lift_total V 0 (i+1)) #W #HVW
-  elim (lift_split … HVW i i) /3 width=7 by cpg_delta, ex2_2_intro/
-| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #l #HLK
-  elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
-  [ elim (IHU1 (L. ⓑ{I} W1) (l+1)) -IHU1 /3 width=9 by cpg_bind, drop_drop, lift_bind, ex2_2_intro/
-  | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpg_flat, lift_flat, ex2_2_intro/
+lemma cpg_lifts: ∀c,h,G. d_liftable2 (cpg h c G).
+#c #h #G #K #T generalize in match c; -c
+@(fqup_wf_ind_eq … G K T) -G -K -T #G0 #K0 #T0 #IH #G #K * *
+[ #s #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct -IH
+  lapply (lifts_inv_sort1 … H1) -H1 #H destruct
+  elim (cpg_inv_sort1 … H2) -H2 * #H1 #H2 destruct
+  /2 width=3 by cpg_atom, cpg_ess, lifts_sort, ex2_intro/
+| #i1 #HG #HK #HT #c #T2 #H2 #b #f #L #HLK #X1 #H1 destruct
+  elim (cpg_inv_lref1_drops … H2) -H2 *
+  [ #H1 #H2 destruct /2 width=3 by ex2_intro/ ]
+  #cV #K0 #V #V2 #HK0 #HV2 #HVT2 #H destruct
+  elim (lifts_inv_lref1 … H1) -H1 #i2 #Hf #H destruct
+  lapply (drops_trans … HLK … HK0 ??) -HLK [3,6: |*: // ] #H
+  elim (drops_split_trans … H) -H [1,6: |*: /2 width=6 by after_uni_dx/ ] #Y #HL0 #HY
+  lapply (drops_inv_tls_at … Hf … HY) -HY #HY
+  elim (drops_inv_skip2 … HY) -HY #L0 #W #HLK0 #HVW #H destruct
+  elim (IH … HV2 … HLK0 … HVW) -IH /2 width=2 by fqup_lref/ -K -K0 -V #W2 #HVW2 #HW2
+  elim (lifts_total W2 (𝐔❴⫯i2❵)) #U2 #HWU2
+  lapply (lifts_trans … HVW2 … HWU2 ??) -HVW2 [3,6: |*: // ] #HVU2
+  lapply (lifts_conf … HVT2 … HVU2 f ?) -V2 [1,3: /2 width=3 by after_uni_succ_sn/ ]
+  /4 width=8 by cpg_ell_drops, cpg_delta_drops, drops_inv_gen, ex2_intro/
+| #l #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct -IH
+  lapply (lifts_inv_gref1 … H1) -H1 #H destruct
+  elim (cpg_inv_gref1 … H2) -H2 #H1 #H2 destruct
+  /2 width=3 by cpg_atom, lifts_gref, ex2_intro/
+| #p #I #V1 #T1 #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct
+  elim (lifts_inv_bind1 … H1) -H1 #W1 #U1 #HVW1 #HTU1 #H destruct
+  elim (cpg_inv_bind1 … H2) -H2 *
+  [ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
+    elim (IH … HV12 … HLK … HVW1) -HV12 //
+    elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    /3 width=5 by cpg_bind, lifts_bind, ex2_intro/
+  | #cT #T2 #HT12 #HXT2 #H1 #H2 #H3 destruct
+    elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ] #U2 #HTU2 #HU12
+    lapply (lifts_trans … HXT2 … HTU2 ??) -T2 [3: |*: // ] #HXU2
+    elim (lifts_split_trans … HXU2 f (𝐔❴⫯O❵)) [2: /2 width=1 by after_uni_one_dx/ ]
+    /3 width=5 by cpg_zeta, ex2_intro/
+  ]
+| #I #V1 #T1 #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct
+  elim (lifts_inv_flat1 … H1) -H1 #W1 #U1 #HVW1 #HTU1 #H destruct
+  elim (cpg_inv_flat1 … H2) -H2 *
+  [ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
+    elim (IH … HV12 … HLK … HVW1) -HV12 -HVW1 //
+    elim (IH … HT12 … HLK … HTU1) -IH -HT12 -HLK -HTU1 //
+    /3 width=5 by cpg_flat, lifts_flat, ex2_intro/
+  | #cT #HT12 #H1 #H2 destruct
+    elim (IH … HT12 … HLK … HTU1) -IH -HT12 -HLK -HTU1 //
+    /3 width=3 by cpg_eps, ex2_intro/
+  | #cV #HV12 #H1 #H2 destruct
+    elim (IH … HV12 … HLK … HVW1) -IH -HV12 -HLK -HVW1 //
+    /3 width=3 by cpg_ee, ex2_intro/
+  | #cV #cY #cT #a #V2 #Y1 #Y2 #T0 #T2 #HV12 #HY12 #HT12 #H1 #H2 #H3 #H4 destruct
+    elim (lifts_inv_bind1 … HTU1) -HTU1 #Z1 #U0 #HYZ1 #HTU1 #H destruct
+    elim (IH … HV12 … HLK … HVW1) -HV12 -HVW1 //
+    elim (IH … HY12 … HLK … HYZ1) -HY12 //
+    elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    /4 width=7 by cpg_beta, lifts_bind, lifts_flat, ex2_intro/
+  | #cV #cY #cT #a #V2 #V20 #Y1 #Y2 #T0 #T2 #HV12 #HV20 #HY12 #HT12 #H1 #H2 #H3 #H4 destruct
+    elim (lifts_inv_bind1 … HTU1) -HTU1 #Z1 #U0 #HYZ1 #HTU1 #H destruct
+    elim (IH … HV12 … HLK … HVW1) -HV12 -HVW1 // #W2 #HVW2 #HW12
+    elim (IH … HY12 … HLK … HYZ1) -HY12 //
+    elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    elim (lifts_total W2 (𝐔❴1❵)) #W20 #HW20
+    lapply (lifts_trans … HVW2 … HW20 ??) -HVW2 [3: |*: // ] #H
+    lapply (lifts_conf … HV20 … H (↑f) ?) -V2 /2 width=3 by after_uni_one_sn/
+    /4 width=9 by cpg_theta, lifts_bind, lifts_flat, ex2_intro/
   ]
 ]
 qed-.
 
-lemma cpg_inv_lref1: ∀h,c,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, c] T2 →
-                     T2 = #i ∨
-                     ∃∃I,K,V,V2. ⬇[i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ➡[h, c] V2 &
-                                 ⬆[O, i+1] V2 ≡ T2.
-#h #c #G #L #T2 #i #H
-elim (cpg_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-[ #s #d #_ #_ #H destruct
-| #I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/
+(* Inversion lemmas with generic slicing for local environments *************)
+
+lemma cpg_inv_lift1: ∀c,h,G. d_deliftable2_sn (cpg h c G).
+#c #h #G #L #U generalize in match c; -c
+@(fqup_wf_ind_eq … G L U) -G -L -U #G0 #L0 #U0 #IH #G #L * *
+[ #s #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct -IH
+  lapply (lifts_inv_sort2 … H1) -H1 #H destruct
+  elim (cpg_inv_sort1 … H2) -H2 * #H1 #H2 destruct
+  /2 width=3 by cpg_atom, cpg_ess, lifts_sort, ex2_intro/
+| #i2 #HG #HL #HU #c #U2 #H2 #b #f #K #HLK #X1 #H1 destruct
+  elim (cpg_inv_lref1_drops … H2) -H2 *
+  [ #H1 #H2 destruct /2 width=3 by ex2_intro/ ]
+  #cW #L0 #W #W2 #HL0 #HW2 #HWU2 #H destruct
+  elim (lifts_inv_lref2 … H1) -H1 #i1 #Hf #H destruct
+  lapply (drops_split_div … HLK (𝐔❴i1❵) ???) -HLK [4,8: * |*: // ] #Y0 #HK0 #HLY0
+  lapply (drops_conf … HL0 … HLY0 ??) -HLY0 [3,6: |*: /2 width=6 by after_uni_dx/ ] #HLY0
+  lapply (drops_inv_tls_at … Hf … HLY0) -HLY0 #HLY0
+  elim (drops_inv_skip1 … HLY0) -HLY0 #K0 #V #HLK0 #HVW #H destruct
+  elim (IH … HW2 … HLK0 … HVW) -IH /2 width=2 by fqup_lref/ -L -L0 -W #V2 #HVW2 #HV2
+  lapply (lifts_trans … HVW2 … HWU2 ??) -W2 [3,6: |*: // ] #HVU2
+  elim (lifts_split_trans … HVU2 ? f) -HVU2 [1,4: |*: /2 width=4 by after_uni_succ_sn/ ]
+  /4 width=8 by cpg_ell_drops, cpg_delta_drops, drops_inv_F, ex2_intro/
+| #l #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct -IH
+  lapply (lifts_inv_gref2 … H1) -H1 #H destruct
+  elim (cpg_inv_gref1 … H2) -H2 #H1 #H2 destruct
+  /2 width=3 by cpg_atom, lifts_gref, ex2_intro/
+| #p #I #W1 #U1 #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct
+  elim (lifts_inv_bind2 … H1) -H1 #V1 #T1 #HVW1 #HTU1 #H destruct
+  elim (cpg_inv_bind1 … H2) -H2 *
+  [ #cW #cU #W2 #U2 #HW12 #HU12 #H1 #H2 destruct
+    elim (IH … HW12 … HLK … HVW1) -HW12 //
+    elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    /3 width=5 by cpg_bind, lifts_bind, ex2_intro/
+  | #cU #U2 #HU12 #HXU2 #H1 #H2 #H3 destruct
+    elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ] #T2 #HTU2 #HT12
+    elim (lifts_div4_one … HTU2 … HXU2) -U2 /3 width=5 by cpg_zeta, ex2_intro/
+  ]
+| #I #W1 #U1 #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct
+  elim (lifts_inv_flat2 … H1) -H1 #V1 #T1 #HVW1 #HTU1 #H destruct
+  elim (cpg_inv_flat1 … H2) -H2 *
+  [ #cW #cU #W2 #U2 #HW12 #HU12 #H1 #H2 destruct
+    elim (IH … HW12 … HLK … HVW1) -HW12 -HVW1 //
+    elim (IH … HU12 … HLK … HTU1) -IH -HU12 -HLK -HTU1 //
+    /3 width=5 by cpg_flat, lifts_flat, ex2_intro/
+  | #cU #HU12 #H1 #H2 destruct
+    elim (IH … HU12 … HLK … HTU1) -IH -HU12 -HLK -HTU1 //
+    /3 width=3 by cpg_eps, ex2_intro/
+  | #cW #HW12 #H1 #H2 destruct
+    elim (IH … HW12 … HLK … HVW1) -IH -HW12 -HLK -HVW1 //
+    /3 width=3 by cpg_ee, ex2_intro/
+  | #cW #cZ #cU #a #W2 #Z1 #Z2 #U0 #U2 #HW12 #HZ12 #HU12 #H1 #H2 #H3 #H4 destruct
+    elim (lifts_inv_bind2 … HTU1) -HTU1 #Y1 #T0 #HYZ1 #HTU1 #H destruct
+    elim (IH … HW12 … HLK … HVW1) -HW12 -HVW1 //
+    elim (IH … HZ12 … HLK … HYZ1) -HZ12 //
+    elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    /4 width=7 by cpg_beta, lifts_bind, lifts_flat, ex2_intro/
+  | #cW #cZ #cU #a #W2 #W20 #Z1 #Z2 #U0 #U2 #HW12 #HW20 #HZ12 #HU12 #H1 #H2 #H3 #H4 destruct
+    elim (lifts_inv_bind2 … HTU1) -HTU1 #Y1 #T0 #HYZ1 #HTU1 #H destruct
+    elim (IH … HW12 … HLK … HVW1) -HW12 -HVW1 // #V2 #HVW2 #HV12
+    elim (IH … HZ12 … HLK … HYZ1) -HZ12 //
+    elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    lapply (lifts_trans … HVW2 … HW20 ??) -W2 [3: |*: // ] #H
+    elim (lifts_split_trans … H ? (↑f)) -H [ |*: /2 width=3 by after_uni_one_sn/ ]
+    /4 width=9 by cpg_theta, lifts_bind, lifts_flat, ex2_intro/
+  ]
 ]
 qed-.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_length.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_length.ma
deleted file mode 100644 (file)
index 142f0fd..0000000
+++ /dev/null
@@ -1,26 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||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 "basic_2/relocation/drops.ma".
-include "basic_2/rt_transition/cpg.ma".
-
-(* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
-
-(* Properties with length for local environments ****************************)
-
-lemma cpg_inv_lref1_ge: ∀h,c,G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, c] T2 → |L| ≤ i → T2 = #i.
-#h #c #G #L #T2 #i #H elim (cpg_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK #_ #_ #HL -h -G -V2 lapply (drop_fwd_length_lt2 … HLK) -K -I -V1
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
-qed-.
index 131d28b9c78a68c14ff96bb153d45023d9423735..f5b94d9050da15f5b2b8e722d9537335900f06a5 100644 (file)
@@ -19,8 +19,8 @@ include "basic_2/rt_transition/cpg.ma".
 
 (* Properties with restricted refinement for local environments *************)
 
-lemma lsubr_cpg_trans: ∀h,c,G. lsub_trans … (cpg h c G) lsubr.
-#h #c #G #L1 #T1 #T2 #H elim H -c -G -L1 -T1 -T2
+lemma lsubr_cpg_trans: ∀c,h,G. lsub_trans … (cpg h c G) lsubr.
+#c #h #G #L1 #T1 #T2 #H elim H -c -G -L1 -T1 -T2
 [ //
 | /2 width=2 by cpg_st/
 | #c #G #L1 #V1 #V2 #W2 #_ #HVW2 #IH #X #H
index 1c9f06d47d80ec440bc652a1db9436cef6abd12f..53448a7bb229e20279a0a2ace105b554d693fb77 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/drops.ma".
+include "basic_2/grammar/term_simple.ma".
 include "basic_2/rt_transition/cpg.ma".
 
 (* CONTEXT-SENSITIVE GENERIC PARALLEL RT-TRANSITION FOR TERMS ***************)
 
-(* Properties with generic slicing for local environments *******************)
+(* Properties with simple terms *********************************************)
 
 (* Note: the main property of simple terms *)
-lemma cpg_inv_appl1_simple: ∀h,c,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, c] U → 𝐒⦃T1⦄ →
-                            ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, c] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, c] T2 &
-                                     U = ⓐV2.T2.
-#h #c #G #L #V1 #T1 #U #H #HT1
-elim (cpg_inv_appl1 … H) -H *
-[ /2 width=5 by ex3_2_intro/
-| #a #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H #_ destruct
+lemma cpg_inv_appl1_simple: ∀c,h,G,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[c, h] U → 𝐒⦃T1⦄ →
+                            ∃∃cV,cT,V2,T2. ⦃G, L⦄ ⊢ V1 ➡[cV, h] V2 & ⦃G, L⦄ ⊢ T1 ➡[cT, h] T2 &
+                                           U = ⓐV2.T2 & c = (↓cV)+cT.
+#c #h #G #L #V1 #T1 #U #H #HT1 elim (cpg_inv_appl1 … H) -H *
+[ /2 width=8 by ex4_4_intro/
+| #cV #cW #cT #p #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #H destruct
   elim (simple_inv_bind … HT1)
-| #a #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H #_ destruct
+| #cV #cW #cT #p #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #H destruct
   elim (simple_inv_bind … HT1)
 ]
 qed-.
index e388af445214426bf245614f35c6b766bfc08d4e..4255a20c4dac165e5fdc3df3d5b61e7ee28f3b74 100644 (file)
@@ -22,7 +22,7 @@ include "basic_2/static/aaa.ma".
 (* Advanced properties ******************************************************)
 
 (* Basic_2A1: was: aaa_lref *)
-lemma aaa_lref_gen: ∀I,G,K,V,B,i,L. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ⁝ B → ⦃G, L⦄ ⊢ #i ⁝ B.
+lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⬇*[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ⁝ B → ⦃G, L⦄ ⊢ #i ⁝ B.
 #I #G #K #V #B #i elim i -i
 [ #L #H lapply (drops_fwd_isid … H ?) -H //
   #H destruct /2 width=1 by aaa_zero/
@@ -34,8 +34,8 @@ qed.
 (* Advanced inversion lemmas ************************************************)
 
 (* Basic_2A1: was: aaa_inv_lref *)
-lemma aaa_inv_lref_gen: ∀G,A,i,L. ⦃G, L⦄ ⊢ #i ⁝ A →
-                        ∃∃I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A.
+lemma aaa_inv_lref_drops: ∀G,A,i,L. ⦃G, L⦄ ⊢ #i ⁝ A →
+                          ∃∃I,K,V. ⬇*[i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ⁝ A.
 #G #A #i elim i -i
 [ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/
 | #i #IH #L #H elim (aaa_inv_lref … H) -H
@@ -53,13 +53,13 @@ lemma aaa_lifts: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀b,f,L2. ⬇*[b, f
   lapply (aaa_inv_sort … H) -H #H destruct
   >(lifts_inv_sort1 … HX) -HX //
 | #i1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
-  elim (aaa_inv_lref_gen … H) -H #J #K1 #V1 #HLK1 #HA
+  elim (aaa_inv_lref_drops … H) -H #J #K1 #V1 #HLK1 #HA
   elim (lifts_inv_lref1 … HX) -HX #i2 #Hf #H destruct
   lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H
   elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK2 #HY
   lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf
   elim (drops_inv_skip2 … HY) -HY #K2 #V2 #HK21 #HV12 #H destruct
-  /4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_gen/
+  /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_gen/
 | #l #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX -b -f -IH
   elim (aaa_inv_gref … H)
 | #p * #V1 #T1 #HG #HL #HT #A #H #b #f #L2 #HL21 #X #HX
@@ -91,13 +91,13 @@ lemma aaa_inv_lifts: ∀G,L2,T2,A. ⦃G, L2⦄ ⊢ T2 ⁝ A → ∀b,f,L1. ⬇*[
   lapply (aaa_inv_sort … H) -H #H destruct
   >(lifts_inv_sort2 … HX) -HX //
 | #i2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
-  elim (aaa_inv_lref_gen … H) -H #J #K2 #V2 #HLK2 #HA
+  elim (aaa_inv_lref_drops … H) -H #J #K2 #V2 #HLK2 #HA
   elim (lifts_inv_lref2 … HX) -HX #i1 #Hf #H destruct
   lapply (drops_split_div … HL21 (𝐔❴i1❵) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY
   lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY
   lapply (drops_inv_tls_at … Hf … HY) -HY #HY -Hf
   elim (drops_inv_skip1 … HY) -HY #K1 #V1 #HK21 #HV12 #H destruct
-  /4 width=12 by aaa_lref_gen, fqup_lref, drops_inv_F/
+  /4 width=12 by aaa_lref_drops, fqup_lref, drops_inv_F/
 | #l #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX -IH -b -f
   elim (aaa_inv_gref … H)
 | #p * #V2 #T2 #HG #HL #HT #A #H #b #f #L1 #HL21 #X #HX
index f8b71a6a5b083d8bbe79a92e8dccbc99113e649a..25bb3c60b8d832f50eae97d178fb93c7183099a8 100644 (file)
@@ -129,6 +129,9 @@ lemma after_inv_const: ∀f1,f2,f,n2,n. n@f1 ⊚ n2@f2 ≡ n@f → f1 ⊚ f2 ≡
 ]
 qed-.
 
+lemma after_inv_total: ∀f1,f2,f. f1 ⊚ f2 ≡ f → f1 ∘ f2 ≗ f.
+/2 width=4 by after_mono/ qed-.
+
 (* Specific forward lemmas **************************************************)
 
 lemma after_fwd_hd: ∀f1,f2,f,n2,n. f1 ⊚ n2@f2 ≡ n@f → f1@❴n2❵ = n.
index 15969a23c46edb255b980c8b2d79c6a40767a72c..a00a10d1696d858896f005f9800cbc240d97b9f1 100644 (file)
@@ -436,6 +436,50 @@ lemma after_uni_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 →
 ]
 qed-.
 
+lemma after_uni_succ_dx: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 →
+                         ∀f. f2 ⊚ 𝐔❴⫯i1❵ ≡ f → 𝐔❴⫯i2❵ ⊚ ⫱*[⫯i2] f2 ≡ f.
+#i2 elim i2 -i2
+[ #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
+  elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H
+  lapply (after_isid_inv_dx … Hg ?) -Hg
+  /4 width=5 by isid_after_sn, after_eq_repl_back_0, after_next/
+| #i2 #IH #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
+  [ #g2 #j1 #Hg2 #H1 #H2 destruct
+    elim (after_inv_pnx … Hf) -Hf [ |*: // ] #g #Hg #H destruct
+    /3 width=5 by after_next/
+  | #g2 #Hg2 #H2 destruct
+    elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
+    /3 width=5 by after_next/
+  ]
+]
+qed.
+
+lemma after_uni_succ_sn: ∀i2,i1,f2. @⦃i1, f2⦄ ≡ i2 →
+                         ∀f. 𝐔❴⫯i2❵ ⊚ ⫱*[⫯i2] f2 ≡ f → f2 ⊚ 𝐔❴⫯i1❵ ≡ f.
+#i2 elim i2 -i2
+[ #i1 #f2 #Hf2 #f #Hf
+  elim (at_inv_xxp … Hf2) -Hf2 // #g2 #H1 #H2 destruct
+  elim (after_inv_nxx … Hf) -Hf [ |*: // ] #g #Hg #H destruct
+  lapply (after_isid_inv_sn … Hg ?) -Hg
+  /4 width=7 by isid_after_dx, after_eq_repl_back_0, after_push/
+| #i2 #IH #i1 #f2 #Hf2 #f #Hf
+  elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H destruct
+  elim (at_inv_xxn … Hf2) -Hf2 [1,3: * |*: // ]
+  [ #g2 #j1 #Hg2 #H1 #H2 destruct <tls_xn in Hg; /3 width=7 by after_push/
+  | #g2 #Hg2 #H2 destruct <tls_xn in Hg; /3 width=5 by after_next/
+  ]
+]
+qed-.
+
+lemma after_uni_one_dx: ∀f2,f. ↑f2 ⊚ 𝐔❴⫯O❵ ≡ f → 𝐔❴⫯O❵ ⊚ f2 ≡ f.
+#f2 #f #H @(after_uni_succ_dx … (↑f2)) /2 width=3 by at_refl/
+qed.
+
+lemma after_uni_one_sn: ∀f1,f. 𝐔❴⫯O❵ ⊚ f1 ≡ f → ↑f1 ⊚ 𝐔❴⫯O❵ ≡ f.
+/3 width=3 by after_uni_succ_sn, at_refl/ qed-.
+
 (* Forward lemmas on istot **************************************************)
 
 lemma after_istot_fwd: ∀f2,f1,f. f2 ⊚ f1 ≡ f → 𝐓⦃f2⦄ → 𝐓⦃f1⦄ → 𝐓⦃f⦄.
index a1f191a2586e98896ca9db07831e3d2793136645..90e73d347f4de952566bf8293abd4e3f38b48d04 100644 (file)
@@ -26,6 +26,10 @@ coinductive at: rtmap → relation nat ≝
 interpretation "relational application (rtmap)"
    'RAt i1 f i2 = (at f i1 i2).
 
+definition H_at_div: relation4 rtmap rtmap rtmap rtmap ≝ λf2,g2,f1,g1.
+                     ∀jf,jg,j. @⦃jf, f2⦄ ≡ j → @⦃jg, g2⦄ ≡ j →
+                     ∃∃j0. @⦃j0, f1⦄ ≡ jf & @⦃j0, g1⦄ ≡ jg.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma at_inv_ppx: ∀f,i1,i2. @⦃i1, f⦄ ≡ i2 → ∀g. 0 = i1 → ↑g = f → 0 = i2.
@@ -267,6 +271,45 @@ theorem at_inj: ∀f,i1,i. @⦃i1, f⦄ ≡ i → ∀i2. @⦃i2, f⦄ ≡ i →
 #Hi elim (lt_le_false i i) /3 width=6 by at_monotonic, eq_sym/
 qed-.
 
+theorem at_div_comm: ∀f2,g2,f1,g1.
+                     H_at_div f2 g2 f1 g1 → H_at_div g2 f2 g1 f1.
+#f2 #g2 #f1 #g1 #IH #jg #jf #j #Hg #Hf
+elim (IH … Hf Hg) -IH -j /2 width=3 by ex2_intro/
+qed-.
+
+theorem at_div_pp: ∀f2,g2,f1,g1.
+                   H_at_div f2 g2 f1 g1 → H_at_div (↑f2) (↑g2) (↑f1) (↑g1).
+#f2 #g2 #f1 #g1 #IH #jf #jg #j #Hf #Hg
+elim (at_inv_xpx … Hf) -Hf [1,2: * |*: // ]
+[ #H1 #H2 destruct -IH
+  lapply (at_inv_xpp … Hg ???) -Hg [4: |*: // ] #H destruct
+  /3 width=3 by at_refl, ex2_intro/
+| #xf #i #Hf2 #H1 #H2 destruct
+  lapply (at_inv_xpn … Hg ????) -Hg [5: * |*: // ] #xg #Hg2 #H destruct
+  elim (IH … Hf2 Hg2) -IH -i /3 width=9 by at_push, ex2_intro/
+]
+qed-.
+
+theorem at_div_nn: ∀f2,g2,f1,g1.
+                   H_at_div f2 g2 f1 g1 → H_at_div (⫯f2) (⫯g2) (f1) (g1).
+#f2 #g2 #f1 #g1 #IH #jf #jg #j #Hf #Hg
+elim (at_inv_xnx … Hf) -Hf [ |*: // ] #i #Hf2 #H destruct
+lapply (at_inv_xnn … Hg ????) -Hg [5: |*: // ] #Hg2
+elim (IH … Hf2 Hg2) -IH -i /2 width=3 by ex2_intro/
+qed-.
+
+theorem at_div_np: ∀f2,g2,f1,g1.
+                   H_at_div f2 g2 f1 g1 → H_at_div (⫯f2) (↑g2) (f1) (⫯g1).
+#f2 #g2 #f1 #g1 #IH #jf #jg #j #Hf #Hg
+elim (at_inv_xnx … Hf) -Hf [ |*: // ] #i #Hf2 #H destruct
+lapply (at_inv_xpn … Hg ????) -Hg [5: * |*: // ] #xg #Hg2 #H destruct
+elim (IH … Hf2 Hg2) -IH -i /3 width=7 by at_next, ex2_intro/
+qed-.
+
+theorem at_div_pn: ∀f2,g2,f1,g1.
+                   H_at_div f2 g2 f1 g1 → H_at_div (↑f2) (⫯g2) (⫯f1) (g1).
+/4 width=6 by at_div_np, at_div_comm/ qed-.
+
 (* Properties on tls ********************************************************)
 
 lemma at_pxx_tls: ∀n,f. @⦃0, f⦄ ≡ n → @⦃0, ⫱*[n]f⦄ ≡ 0.
@@ -321,6 +364,22 @@ lemma id_inv_at: ∀f. (∀i. @⦃i, f⦄ ≡ i) → 𝐈𝐝 ≗ f.
 lemma id_at: ∀i. @⦃i, 𝐈𝐝⦄ ≡ i.
 /2 width=1 by isid_inv_at/ qed.
 
+(* Advanced forward lemmas on id ********************************************)
+
+lemma at_id_fwd: ∀i1,i2. @⦃i1, 𝐈𝐝⦄ ≡ i2 → i1 = i2.
+/2 width=4 by at_mono/ qed.
+
+(* Main properties on id ****************************************************)
+
+theorem at_div_id_dx: ∀f. H_at_div f 𝐈𝐝 𝐈𝐝 f.
+#f #jf #j0 #j #Hf #H0
+lapply (at_id_fwd … H0) -H0 #H destruct
+/2 width=3 by ex2_intro/
+qed-.
+
+theorem at_div_id_sn: ∀f. H_at_div 𝐈𝐝 f f 𝐈𝐝.
+/3 width=6 by at_div_id_dx, at_div_comm/ qed-.
+
 (* Properties with uniform relocations **************************************)
 
 lemma at_uni: ∀n,i. @⦃i,𝐔❴n❵⦄ ≡ n+i.