]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / frees_drops.ma
index 3c64660863240b8654d66b9f8aa03e58b7815299..37c629351f1c30b64f19adb8ee78124871942f2a 100644 (file)
@@ -20,8 +20,9 @@ include "static_2/static/frees_fqup.ma".
 
 (* Advanced properties ******************************************************)
 
-lemma frees_atom_drops: โˆ€b,L,i. โฌ‡*[b, ๐”โดiโต] L โ‰˜ โ‹† โ†’
-                        โˆ€f. ๐ˆโฆƒfโฆ„ โ†’ L โŠข ๐…*โฆƒ#iโฆ„ โ‰˜ โซฏ*[i]โ†‘f.
+lemma frees_atom_drops:
+      โˆ€b,L,i. โ‡ฉ*[b,๐”โจiโฉ] L โ‰˜ โ‹† โ†’
+      โˆ€f. ๐ˆโชfโซ โ†’ L โŠข ๐…+โช#iโซ โ‰˜ โซฏ*[i]โ†‘f.
 #b #L elim L -L /2 width=1 by frees_atom/
 #L #I #IH *
 [ #H lapply (drops_fwd_isid โ€ฆ H ?) -H // #H destruct
@@ -29,55 +30,42 @@ lemma frees_atom_drops: โˆ€b,L,i. โฌ‡*[b, ๐”โดiโต] L โ‰˜ โ‹† โ†’
 ]
 qed.
 
-lemma frees_pair_drops: โˆ€f,K,V. K โŠข ๐…*โฆƒVโฆ„ โ‰˜ f โ†’ 
-                        โˆ€i,I,L. โฌ‡*[i] L โ‰˜ K.โ“‘{I}V โ†’ L โŠข ๐…*โฆƒ#iโฆ„ โ‰˜ โซฏ*[i] โ†‘f.
+lemma frees_pair_drops:
+      โˆ€f,K,V. K โŠข ๐…+โชVโซ โ‰˜ f โ†’
+      โˆ€i,I,L. โ‡ฉ*[i] L โ‰˜ K.โ“‘[I]V โ†’ L โŠข ๐…+โช#iโซ โ‰˜ โซฏ*[i] โ†‘f.
 #f #K #V #Hf #i elim i -i
 [ #I #L #H lapply (drops_fwd_isid โ€ฆ H ?) -H /2 width=1 by frees_pair/
 | #i #IH #I #L #H elim (drops_inv_succ โ€ฆ H) -H /3 width=2 by frees_lref/
 ]
 qed.
 
-lemma frees_unit_drops: โˆ€f.  ๐ˆโฆƒfโฆ„ โ†’ โˆ€I,K,i,L. โฌ‡*[i] L โ‰˜ K.โ“ค{I} โ†’
-                       L โŠข ๐…*โฆƒ#iโฆ„ โ‰˜ โซฏ*[i] โ†‘f.
+lemma frees_unit_drops:
+      โˆ€f.  ๐ˆโชfโซ โ†’ โˆ€I,K,i,L. โ‡ฉ*[i] L โ‰˜ K.โ“ค[I] โ†’
+      L โŠข ๐…+โช#iโซ โ‰˜ โซฏ*[i] โ†‘f.
 #f #Hf #I #K #i elim i -i
 [ #L #H lapply (drops_fwd_isid โ€ฆ H ?) -H /2 width=1 by frees_unit/
 | #i #IH #Y #H elim (drops_inv_succ โ€ฆ H) -H
   #J #L #HLK #H destruct /3 width=1 by frees_lref/
 ]
 qed.
-(*
-lemma frees_sort_pushs: โˆ€f,K,s. K โŠข ๐…*โฆƒโ‹†sโฆ„ โ‰˜ f โ†’
-                        โˆ€i,L. โฌ‡*[i] L โ‰˜ K โ†’ L โŠข ๐…*โฆƒโ‹†sโฆ„ โ‰˜ โซฏ*[i] f.
-#f #K #s #Hf #i elim i -i
-[ #L #H lapply (drops_fwd_isid โ€ฆ H ?) -H //
-| #i #IH #L #H elim (drops_inv_succ โ€ฆ H) -H /3 width=1 by frees_sort/
-]
-qed.
-*)
-lemma frees_lref_pushs: โˆ€f,K,j. K โŠข ๐…*โฆƒ#jโฆ„ โ‰˜ f โ†’
-                        โˆ€i,L. โฌ‡*[i] L โ‰˜ K โ†’ L โŠข ๐…*โฆƒ#(i+j)โฆ„ โ‰˜ โซฏ*[i] f.
+
+lemma frees_lref_pushs:
+      โˆ€f,K,j. K โŠข ๐…+โช#jโซ โ‰˜ f โ†’
+      โˆ€i,L. โ‡ฉ*[i] L โ‰˜ K โ†’ L โŠข ๐…+โช#(i+j)โซ โ‰˜ โซฏ*[i] f.
 #f #K #j #Hf #i elim i -i
 [ #L #H lapply (drops_fwd_isid โ€ฆ H ?) -H //
 | #i #IH #L #H elim (drops_inv_succ โ€ฆ H) -H
   #I #Y #HYK #H destruct /3 width=1 by frees_lref/
 ]
 qed.
-(*
-lemma frees_gref_pushs: โˆ€f,K,l. K โŠข ๐…*โฆƒยงlโฆ„ โ‰˜ f โ†’
-                        โˆ€i,L. โฌ‡*[i] L โ‰˜ K โ†’ L โŠข ๐…*โฆƒยงlโฆ„ โ‰˜ โซฏ*[i] f.
-#f #K #l #Hf #i elim i -i
-[ #L #H lapply (drops_fwd_isid โ€ฆ H ?) -H //
-| #i #IH #L #H elim (drops_inv_succ โ€ฆ H) -H /3 width=1 by frees_gref/
-]
-qed.
-*)
+
 (* Advanced inversion lemmas ************************************************)
 
-lemma frees_inv_lref_drops: โˆ€L,i,f. L โŠข ๐…*โฆƒ#iโฆ„ โ‰˜ f โ†’
-                            โˆจโˆจ โˆƒโˆƒg. โฌ‡*[โ’ป, ๐”โดiโต] L โ‰˜ โ‹† & ๐ˆโฆƒgโฆ„ & f = โซฏ*[i] โ†‘g
-                             | โˆƒโˆƒg,I,K,V. K โŠข ๐…*โฆƒVโฆ„ โ‰˜ g &
-                                          โฌ‡*[i] L โ‰˜ K.โ“‘{I}V & f = โซฏ*[i] โ†‘g
-                             | โˆƒโˆƒg,I,K. โฌ‡*[i] L โ‰˜ K.โ“ค{I} & ๐ˆโฆƒgโฆ„ & f = โซฏ*[i] โ†‘g.
+lemma frees_inv_lref_drops:
+      โˆ€L,i,f. L โŠข ๐…+โช#iโซ โ‰˜ f โ†’
+      โˆจโˆจ โˆƒโˆƒg. โ‡ฉ*[โ’ป,๐”โจiโฉ] L โ‰˜ โ‹† & ๐ˆโชgโซ & f = โซฏ*[i] โ†‘g
+       | โˆƒโˆƒg,I,K,V. K โŠข ๐…+โชVโซ โ‰˜ g & โ‡ฉ*[i] L โ‰˜ K.โ“‘[I]V & f = โซฏ*[i] โ†‘g
+       | โˆƒโˆƒg,I,K. โ‡ฉ*[i] L โ‰˜ K.โ“ค[I] & ๐ˆโชgโซ & f = โซฏ*[i] โ†‘g.
 #L elim L -L
 [ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H
 [ elim (frees_inv_atom โ€ฆ H) -H #f #Hf #H destruct
@@ -97,9 +85,10 @@ qed-.
 
 (* Properties with generic slicing for local environments *******************)
 
-lemma frees_lifts: โˆ€b,f1,K,T. K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f1 โ†’
-                   โˆ€f,L. โฌ‡*[b, f] L โ‰˜ K โ†’ โˆ€U. โฌ†*[f] T โ‰˜ U โ†’
-                   โˆ€f2. f ~โŠš f1 โ‰˜ f2 โ†’ L โŠข ๐…*โฆƒUโฆ„ โ‰˜ f2.
+lemma frees_lifts:
+      โˆ€b,f1,K,T. K โŠข ๐…+โชTโซ โ‰˜ f1 โ†’
+      โˆ€f,L. โ‡ฉ*[b,f] L โ‰˜ K โ†’ โˆ€U. โ‡ง*[f] T โ‰˜ U โ†’
+      โˆ€f2. f ~โŠš f1 โ‰˜ f2 โ†’ L โŠข ๐…+โชUโซ โ‰˜ f2.
 #b #f1 #K #T #H lapply (frees_fwd_isfin โ€ฆ H) elim H -f1 -K -T
 [ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
   lapply (coafter_isid_inv_dx โ€ฆ H3 โ€ฆ Hf1) -f1 #Hf2
@@ -154,48 +143,54 @@ lemma frees_lifts: โˆ€b,f1,K,T. K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f1 โ†’
 ]
 qed-.
 
-lemma frees_lifts_SO: โˆ€b,L,K. โฌ‡*[b, ๐”โด1โต] L โ‰˜ K โ†’ โˆ€T,U. โฌ†*[1] T โ‰˜ U โ†’
-                      โˆ€f. K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f โ†’ L โŠข ๐…*โฆƒUโฆ„ โ‰˜ โซฏf.
+lemma frees_lifts_SO:
+      โˆ€b,L,K. โ‡ฉ*[b,๐”โจ1โฉ] L โ‰˜ K โ†’ โˆ€T,U. โ‡ง*[1] T โ‰˜ U โ†’
+      โˆ€f. K โŠข ๐…+โชTโซ โ‰˜ f โ†’ L โŠข ๐…+โชUโซ โ‰˜ โซฏf.
 #b #L #K #HLK #T #U #HTU #f #Hf
 @(frees_lifts b โ€ฆ Hf โ€ฆ HTU) //  (**) (* auto fails *)
 qed.
 
 (* Forward lemmas with generic slicing for local environments ***************)
 
-lemma frees_fwd_coafter: โˆ€b,f2,L,U. L โŠข ๐…*โฆƒUโฆ„ โ‰˜ f2 โ†’
-                         โˆ€f,K. โฌ‡*[b, f] L โ‰˜ K โ†’ โˆ€T. โฌ†*[f] T โ‰˜ U โ†’
-                         โˆ€f1. K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f1 โ†’ f ~โŠš f1 โ‰˜ f2.
+lemma frees_fwd_coafter:
+      โˆ€b,f2,L,U. L โŠข ๐…+โชUโซ โ‰˜ f2 โ†’
+      โˆ€f,K. โ‡ฉ*[b,f] L โ‰˜ K โ†’ โˆ€T. โ‡ง*[f] T โ‰˜ U โ†’
+      โˆ€f1. K โŠข ๐…+โชTโซ โ‰˜ f1 โ†’ f ~โŠš f1 โ‰˜ f2.
 /4 width=11 by frees_lifts, frees_mono, coafter_eq_repl_back0/ qed-.
 
 (* Inversion lemmas with generic slicing for local environments *************)
 
-lemma frees_inv_lifts_ex: โˆ€b,f2,L,U. L โŠข ๐…*โฆƒUโฆ„ โ‰˜ f2 โ†’
-                          โˆ€f,K. โฌ‡*[b, f] L โ‰˜ K โ†’ โˆ€T. โฌ†*[f] T โ‰˜ U โ†’
-                          โˆƒโˆƒf1. f ~โŠš f1 โ‰˜ f2 & K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f1.
+lemma frees_inv_lifts_ex:
+      โˆ€b,f2,L,U. L โŠข ๐…+โชUโซ โ‰˜ f2 โ†’
+      โˆ€f,K. โ‡ฉ*[b,f] L โ‰˜ K โ†’ โˆ€T. โ‡ง*[f] T โ‰˜ U โ†’
+      โˆƒโˆƒf1. f ~โŠš f1 โ‰˜ f2 & K โŠข ๐…+โชTโซ โ‰˜ f1.
 #b #f2 #L #U #Hf2 #f #K #HLK #T elim (frees_total K T)
 /3 width=9 by frees_fwd_coafter, ex2_intro/
 qed-.
 
-lemma frees_inv_lifts_SO: โˆ€b,f,L,U. L โŠข ๐…*โฆƒUโฆ„ โ‰˜ f โ†’
-                          โˆ€K. โฌ‡*[b, ๐”โด1โต] L โ‰˜ K โ†’ โˆ€T. โฌ†*[1] T โ‰˜ U โ†’
-                          K โŠข ๐…*โฆƒTโฆ„ โ‰˜ โซฑf.
+lemma frees_inv_lifts_SO:
+      โˆ€b,f,L,U. L โŠข ๐…+โชUโซ โ‰˜ f โ†’
+      โˆ€K. โ‡ฉ*[b,๐”โจ1โฉ] L โ‰˜ K โ†’ โˆ€T. โ‡ง*[1] T โ‰˜ U โ†’
+      K โŠข ๐…+โชTโซ โ‰˜ โซฑf.
 #b #f #L #U #H #K #HLK #T #HTU elim(frees_inv_lifts_ex โ€ฆ H โ€ฆ HLK โ€ฆ HTU) -b -L -U
 #f1 #Hf #Hf1 elim (coafter_inv_nxx โ€ฆ Hf) -Hf
 /3 width=5 by frees_eq_repl_back, coafter_isid_inv_sn/
 qed-.
 
-lemma frees_inv_lifts: โˆ€b,f2,L,U. L โŠข ๐…*โฆƒUโฆ„ โ‰˜ f2 โ†’
-                       โˆ€f,K. โฌ‡*[b, f] L โ‰˜ K โ†’ โˆ€T. โฌ†*[f] T โ‰˜ U โ†’
-                       โˆ€f1. f ~โŠš f1 โ‰˜ f2 โ†’ K โŠข ๐…*โฆƒTโฆ„ โ‰˜ f1.
+lemma frees_inv_lifts:
+      โˆ€b,f2,L,U. L โŠข ๐…+โชUโซ โ‰˜ f2 โ†’
+      โˆ€f,K. โ‡ฉ*[b,f] L โ‰˜ K โ†’ โˆ€T. โ‡ง*[f] T โ‰˜ U โ†’
+      โˆ€f1. f ~โŠš f1 โ‰˜ f2 โ†’ K โŠข ๐…+โชTโซ โ‰˜ f1.
 #b #f2 #L #U #H #f #K #HLK #T #HTU #f1 #Hf2 elim (frees_inv_lifts_ex โ€ฆ H โ€ฆ HLK โ€ฆ HTU) -b -L -U
 /3 width=7 by frees_eq_repl_back, coafter_inj/
 qed-.
 
 (* Note: this is used by rex_conf and might be modified *)
-lemma frees_inv_drops_next: โˆ€f1,L1,T1. L1 โŠข ๐…*โฆƒT1โฆ„ โ‰˜ f1 โ†’
-                            โˆ€I2,L2,V2,n. โฌ‡*[n] L1 โ‰˜ L2.โ“‘{I2}V2 โ†’
-                            โˆ€g1. โ†‘g1 = โซฑ*[n] f1 โ†’
-                            โˆƒโˆƒg2. L2 โŠข ๐…*โฆƒV2โฆ„ โ‰˜ g2 & g2 โŠ† g1.
+lemma frees_inv_drops_next:
+      โˆ€f1,L1,T1. L1 โŠข ๐…+โชT1โซ โ‰˜ f1 โ†’
+      โˆ€I2,L2,V2,n. โ‡ฉ*[n] L1 โ‰˜ L2.โ“‘[I2]V2 โ†’
+      โˆ€g1. โ†‘g1 = โซฑ*[n] f1 โ†’
+      โˆƒโˆƒg2. L2 โŠข ๐…+โชV2โซ โ‰˜ g2 & g2 โŠ† g1.
 #f1 #L1 #T1 #H elim H -f1 -L1 -T1
 [ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s
   lapply (isid_tls n โ€ฆ Hf1) -Hf1 <H1 -f1 #Hf1