]> matita.cs.unibo.it Git - helm.git/commitdiff
some advances on pointwise union for local environments ...
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 8 Apr 2014 20:56:32 +0000 (20:56 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 8 Apr 2014 20:56:32 +0000 (20:56 +0000)
matita/matita/contribs/lambdadelta/basic_2/relocation/llor.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/lpx_sn_alt.ma
matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl

index 8c76ecc02f252dfa46269bde8184dfc100c63800..7dec554c7be07008a8cec2cf593bbe7b7c23069c 100644 (file)
@@ -15,7 +15,7 @@
 include "basic_2/notation/relations/lazyor_4.ma".
 include "basic_2/relocation/lpx_sn_alt.ma".
 
-(* LAZY UNION FOR LOCAL ENVIRONMENTS ****************************************)
+(* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************)
 
 inductive clor (T) (L2) (K1) (V1): predicate term ≝
 | clor_sn: ∀U. |K1| < |L2| → ⇧[|L2|-|K1|-1, 1] U ≡ T → clor T L2 K1 V1 V1
@@ -53,3 +53,48 @@ lemma llor_total: ∀T,L2,L1. |L1| ≤ |L2| → ∃L. L1 ⩖[T] L2 ≡ L.
   /5 width=4 by llor_pair_dx, monotonic_lt_minus_l, ex_intro/
 ]
 qed-.
+
+(* Alternative definition ***************************************************)
+
+lemma llor_intro_alt_eq: ∀T,L2,L1,L. |L1| = |L2| → |L1| = |L| →
+                         (∀I1,I,K1,K,V1,V,U,i. ⇧[i, 1] U ≡ T →
+                            ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L ≡ K.ⓑ{I}V →
+                            ∧∧ I1 = I & V1 = V & K1 ⩖[T] L2 ≡ K
+                         ) →
+                         (∀I1,I2,I,K1,K2,K,V1,V2,V,i. (∀U. ⇧[i, 1] U ≡ T → ⊥) →
+                            ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
+                            ⇩[i] L ≡ K.ⓑ{I}V → ∧∧ I1 = I & V2 = V & K1 ⩖[T] L2 ≡ K
+                         ) → L1 ⩖[T] L2 ≡ L.
+#T #L2 #L1 #L #HL12 #HL1 #IH1 #IH2 @lpx_sn_intro_alt // -HL1
+#I1 #I #K1 #K #V1 #V #i #HLK1 #HLK
+lapply (ldrop_fwd_length_minus4 … HLK1)
+lapply (ldrop_fwd_length_le4 … HLK1)
+normalize >HL12 <minus_plus #HKL1 #Hi elim (is_lift_dec T i 1) -HL12
+[ * #U #HUT elim (IH1 … HUT HLK1 HLK) -IH1 -HLK1 -HLK #H1 #H2 #HT destruct
+  /3 width=2 by clor_sn, and3_intro/
+| #H elim (ldrop_O1_lt L2 i) destruct /2 width=1 by monotonic_lt_minus_l/
+  #I2 #K2 #V2 #HLK2 elim (IH2 … HLK1 HLK2 HLK) -HLK1 -HLK
+  /5 width=3 by clor_dx, ex_intro, and3_intro/
+]
+qed.
+
+lemma llor_inv_gen_eq: ∀T,L2,L1,L. L1 ⩖[T] L2 ≡ L → |L1| = |L2| →
+                       |L1| = |L| ∧
+                       (∀I1,I,K1,K,V1,V,i.
+                          ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L ≡ K.ⓑ{I}V →
+                          (∃∃U. ⇧[i, 1] U ≡ T &
+                                I1 = I & V1 = V & K1 ⩖[T] L2 ≡ K
+                          ) ∨
+                          (∃∃I2,K2,V2. (∀U. ⇧[i, 1] U ≡ T → ⊥) &
+                                       ⇩[i] L2 ≡ K2.ⓑ{I2}V2 &
+                                       I1 = I & V2 = V & K1 ⩖[T] L2 ≡ K
+                          )
+                       ).
+#T #L2 #L1 #L #H #HL12 elim (lpx_sn_inv_gen … H) -H
+#HL1 #IH @conj //
+#I1 #I #K1 #K #V1 #V #i #HLK1 #HLK
+lapply (ldrop_fwd_length_minus4 … HLK1)
+lapply (ldrop_fwd_length_le4 … HLK1)
+normalize >HL12 <minus_plus #HKL1 #Hi elim (IH … HLK1 HLK) -IH #H *
+/4 width=5 by ex5_3_intro, ex4_intro, or_intror, or_introl/
+qed-.
index b5aa43cf8e6279f3d97c95e2e9c975519ae59a35..c9c381a32b300fc7456b6736ea4d4a662216450b 100644 (file)
@@ -137,7 +137,7 @@ qed.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma lpx_sn_inv_alt: ∀R,L1,L2. lpx_sn R L1 L2 →
+lemma lpx_sn_inv_gen: ∀R,L1,L2. lpx_sn R L1 L2 →
                       |L1| = |L2| ∧
                       ∀I1,I2,K1,K2,V1,V2,i.
                       ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 →
index f65ec6eee8d79675e0337967c4585c4551baf2a0..e274719b76267a5c8b023a365b16a27c295b389f 100644 (file)
@@ -247,12 +247,17 @@ table {
              [ "gdrop ( ⇩[?] ? ≡ ? )" "gdrop_gdrop" * ]
           }
         ]
-        [ { "lazy pointwise extension of a relation" * } {
+        [ { "pointwise union for local environments" * } {
+             [ "llor ( ? ⩖[?] ? ≡ ? )" * ]
+          }
+        ]
+        [ { "pointwise extension of a relation" * } {
              [ "llpx_sn" "llpx_sn_alt" + "llpx_sn_tc" + "llpx_sn_leq" + "llpx_sn_ldrop" + "llpx_sn_lpx_sn" * ]
+             [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_ldrop" + "lpx_sn_lpx_sn" * ]
           }
         ]
         [ { "basic local env. slicing" * } {
-             [ "ldrop ( ⇩[?,?,?] ? ≡ ? )" "ldrop_lpx_sn" + "ldrop_leq" + "ldrop_ldrop" * ]
+             [ "ldrop ( ⇩[?,?,?] ? ≡ ? )"  "ldrop_leq" + "ldrop_ldrop" * ]
           }
         ]
         [ { "basic term relocation" * } {
@@ -268,10 +273,6 @@ table {
              [ "leq ( ? ≃[?,?] ? )" "leq_leq" * ]
           }
         ]
-        [ { "pointwise extension of a relation" * } {
-             [ "lpx_sn" "lpx_sn_alt" + "lpx_sn_tc" + "lpx_sn_lpx_sn" * ]
-          }
-        ]
         [ { "same top term constructor" * } {
              [ "tstc ( ? ≃ ? )" "tstc_tstc" + "tstc_vector" * ]
           }