From: Ferruccio Guidi <ferruccio.guidi@unibo.it>
Date: Tue, 30 Sep 2014 14:36:42 +0000 (+0000)
Subject: notation update for pointwise union
X-Git-Tag: make_still_working~833
X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=c7c0bdef0ba026f4ba85528cba8cb33f6cf2ef5d;p=helm.git

notation update for pointwise union
---

diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/lleq_llor.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/lleq_llor.ma
index df65d267a..1cfaa4013 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/multiple/lleq_llor.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/lleq_llor.ma
@@ -21,7 +21,7 @@ include "basic_2/multiple/lleq_alt.ma".
 (* Properties on pointwise union for local environments **********************)
 
 lemma llpx_sn_llor_dx: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
-                       ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⩖[T, d] L2 ≡ L → L2 ≡[T, d] L.
+                       ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⋓[T, d] L2 ≡ L → L2 ≡[T, d] L.
 #R #H1R #H2R #L1 #L2 #T #d #H1 #L #H2
 lapply (llpx_sn_frees_trans … H1R H2R … H1) -H1R -H2R #HR
 elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
@@ -37,5 +37,5 @@ elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK1 -HLK2 -HLK * /2 width=1 by conj/ #H
 qed.
 
 lemma llpx_sn_llor_dx_sym: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) →
-                           ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⩖[T, d] L2 ≡ L → L ≡[T, d] L2.
+                           ∀L1,L2,T,d. llpx_sn R d T L1 L2 → ∀L. L1 ⋓[T, d] L2 ≡ L → L ≡[T, d] L2.
 /3 width=6 by llpx_sn_llor_dx, lleq_sym/ qed.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma
index 1e124cdb5..29d4c1d51 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma
@@ -33,7 +33,7 @@ interpretation
 (* Basic properties *********************************************************)
 
 (* Note: this can be proved by llor_skip *)
-lemma llor_atom: ∀T,d. ⋆ ⩖[T, d] ⋆ ≡ ⋆.
+lemma llor_atom: ∀T,d. ⋆ ⋓[T, d] ⋆ ≡ ⋆.
 #T #d @and3_intro //
 #I1 #I2 #I #K1 #K2 #K #V1 #V2 #V #i #HLK1
 elim (drop_inv_atom1 … HLK1) -HLK1 #H destruct
diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_alt.ma
index 0b634cbf9..721d18c94 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_alt.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_alt.ma
@@ -19,9 +19,9 @@ include "basic_2/multiple/llor.ma".
 
 (* Alternative definition ***************************************************)
 
-lemma llor_tail_frees: ∀L1,L2,L,U,d. L1 ⩖[U, d] L2 ≡ L → d ≤ yinj (|L1|) →
+lemma llor_tail_frees: ∀L1,L2,L,U,d. L1 ⋓[U, d] L2 ≡ L → d ≤ yinj (|L1|) →
                        ∀I1,W1. ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ →
-                       ∀I2,W2. ⓑ{I1}W1.L1 ⩖[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L.
+                       ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L.
 #L1 #L2 #L #U #d * #HL12 #HL1 #IH #Hd #I1 #W1 #HU #I2 #W2
 @and3_intro [1,2: >ltail_length /2 width=1 by le_S_S/ ]
 #J1 #J2 #J #K1 #K2 #K #V1 #V2 #V #i #HLK1 #HLK2 #HLK
@@ -46,9 +46,9 @@ elim (le_to_or_lt_eq … H) -H #H
 ]
 qed.
 
-lemma llor_tail_cofrees: ∀L1,L2,L,U,d. L1 ⩖[U, d] L2 ≡ L →
+lemma llor_tail_cofrees: ∀L1,L2,L,U,d. L1 ⋓[U, d] L2 ≡ L →
                          ∀I1,W1. (ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ → ⊥) →
-                         ∀I2,W2. ⓑ{I1}W1.L1 ⩖[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
+                         ∀I2,W2. ⓑ{I1}W1.L1 ⋓[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
 #L1 #L2 #L #U #d * #HL12 #HL1 #IH #I1 #W1 #HU #I2 #W2
 @and3_intro [1,2: >ltail_length /2 width=1 by le_S_S/ ]
 #J1 #J2 #J #K1 #K2 #K #V1 #V2 #V #i #HLK1 #HLK2 #HLK
diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_drop.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_drop.ma
index d839edcc7..8692a91f1 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_drop.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_drop.ma
@@ -19,7 +19,7 @@ include "basic_2/multiple/llor_alt.ma".
 
 (* Advanced properties ******************************************************)
 
-lemma llor_skip: ∀L1,L2,U,d. |L1| = |L2| → yinj (|L1|) ≤ d → L1 ⩖[U, d] L2 ≡ L1.
+lemma llor_skip: ∀L1,L2,U,d. |L1| = |L2| → yinj (|L1|) ≤ d → L1 ⋓[U, d] L2 ≡ L1.
 #L1 #L2 #U #d #HL12 #Hd @and3_intro // -HL12
 #I1 #I2 #I #K1 #K2 #K #W1 #W2 #W #i #HLK1 #_ #HLK -L2 -K2
 lapply (drop_mono … HLK … HLK1) -HLK #H destruct
@@ -28,7 +28,7 @@ lapply (drop_fwd_length_lt2 … HLK1) -K1
 qed.
 
 (* Note: lemma 1400 concludes the "big tree" theorem *)
-lemma llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⩖[T, d] L2 ≡ L.
+lemma llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⋓[T, d] L2 ≡ L.
 #L1 @(lenv_ind_alt … L1) -L1
 [ #L2 #T #d #H >(length_inv_zero_sn … H) -L2 /2 width=2 by ex_intro/
 | #I1 #L1 #V1 #IHL1 #Y #T #d >ltail_length #H
diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma
index e5933d4c7..91d9e649e 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma
@@ -22,7 +22,7 @@ include "basic_2/multiple/lleq_alt.ma".
 
 lemma llpx_sn_llor_fwd_sn: ∀R. (∀L. reflexive … (R L)) →
                            ∀L1,L2,T,d. llpx_sn R d T L1 L2 →
-                           ∀L. L1 ⩖[T, d] L2 ≡ L → lpx_sn R L1 L.
+                           ∀L. L1 ⋓[T, d] L2 ≡ L → lpx_sn R L1 L.
 #R #HR #L1 #L2 #T #d #H1 #L #H2
 elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
 elim H2 -H2 #_ #HL1 #IH2
diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyor_5.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyor_5.ma
index 509089ddb..905905214 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyor_5.ma
+++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyor_5.ma
@@ -14,6 +14,6 @@
 
 (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
 
-notation "hvbox( L1 ⩖ break [ term 46 T , break term 46 d ] break term 46 L2 ≡ break term 46 L )"
+notation "hvbox( L1 ⋓ break [ term 46 T , break term 46 d ] break term 46 L2 ≡ break term 46 L )"
    non associative with precedence 45
    for @{ 'LazyOr $L1 $T $d $L2 $L }.
diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl
index 4b1bbe19f..084720512 100644
--- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl
+++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl
@@ -223,7 +223,7 @@ table {
           }
         ]
         [ { "pointwise union for local environments" * } {
-             [ "llor ( ? ⩖[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ]
+             [ "llor ( ? ⋓[?,?] ? ≡ ? )" "llor_alt" + "llor_drop" * ]
           }
         ]
         [ { "context-sensitive exclusion from free variables" * } {