]> matita.cs.unibo.it Git - helm.git/commitdiff
general slicing reactivated ...
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 9 Feb 2016 19:05:48 +0000 (19:05 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 9 Feb 2016 19:05:48 +0000 (19:05 +0000)
matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc [deleted file]
matita/matita/contribs/lambdadelta/basic_2/names.txt
matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma
matita/matita/contribs/lambdadelta/basic_2/relocation/drops_weight.ma
matita/matita/predefined_virtuals.ml

diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc
deleted file mode 100644 (file)
index ebb4a24..0000000
+++ /dev/null
@@ -1,29 +0,0 @@
-(* Basic_1: was: lift_r *)
-lemma lift_refl: ∀T,l. ⬆[l, 0] T ≡ T.
-#T elim T -T
-[ * #i // #l elim (lt_or_ge i l) /2 width=1 by lift_lref_lt,
-lift_lref_ge/
-| * /2 width=1 by lift_bind, lift_flat/
-]
-qed.
-
-lemma lift_total: ∀T1,l,m. ∃T2. ⬆[l,m] T1 ≡ T2.
-#T1 elim T1 -T1
-[ * #i /2 width=2 by lift_sort, lift_gref, ex_intro/
-  #l #m elim (lt_or_ge i l) /3 width=2 by lift_lref_lt, lift_lref_ge,
-ex_intro/
-| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #l #m
-  elim (IHV1 l m) -IHV1 #V2 #HV12
-  [ elim (IHT1 (l+1) m) -IHT1 /3 width=2 by lift_bind, ex_intro/
-  | elim (IHT1 l m) -IHT1 /3 width=2 by lift_flat, ex_intro/
-  ]
-]
-qed.
-
-lemma liftv_total: ∀l,m. ∀T1s:list term. ∃T2s. ⬆[l, m] T1s ≡ T2s.
-#l #m #T1s elim T1s -T1s
-[ /2 width=2 by liftv_nil, ex_intro/
-| #T1 #T1s * #T2s #HT12s
-  elim (lift_total T1 l m) /3 width=2 by liftv_cons, ex_intro/
-]
-qed-.
index ed799bf447c78672a04de40700cd8cf7f15ae641..1a26839d59bfc247a9fbe8e8ef3092597f7a80d5 100644 (file)
@@ -15,24 +15,22 @@ S      : RTM stack
 T,U,V,W: term
 X,Y,Z  : reserved: transient objet denoted by a capital letter
 
-a,b    : binder polarity
-c      : relocation
+a,b    :
+c      : local dropping kind parameter (true = restricted, false = general)
 d      : term degree
 e      : reserved: future use (\lambda\delta 3)
-f      :
-g      : sort degree parameter
-h      : sort hierarchy parameter
-i,j    : local reference position index (de Bruijn's)
-k      : sort index
-l      : relocation depth
-m      : relocation height
+f,g    : local reference transforming map 
+h      : sort degree parameter
+i,j    : local reference depth (de Bruijn's)
+k,l    : global reference level
+m      : 
 n      : type iterations
-o      :
-p,q    : global reference position
+o      : sort hierarchy parameter
+p,q    : binder polarity
 r      : reduction kind parameter (true = ordinary, false = extended)
-s      : local dropping kind parameter (true = general, false = restricted)
-t,u    : local reference position level (de Bruijn's) (RTM)
-v,w    :
+s      : sort index
+t,u    :
+v,w    : local reference position level (de Bruijn's) (RTM)
 x,y,z  : reserved: transient objet denoted by a small letter
 
 NAMING CONVENTIONS FOR CONSTRUCTORS
@@ -65,6 +63,7 @@ t: context-free for terms
 
 - second letter
 
+e: reserved for generic entrywise extension
 i: irreducible form
 n: normal form
 p: reflexive parallel transformation
@@ -96,4 +95,3 @@ q: reflexive closure                           (question)
 r: proper multiple step (restricted)           (restricted)
 s: reflexive transitive closure                (star)
 u: proper single step (restricted)             (unit)
-x: reserved for generic pointwise extension
index 1016ec0d5c050837a8e11fc8786697d6ff29e5c6..a38d5c7a56d765a68f573e1b5066dc193d9424a7 100644 (file)
@@ -22,165 +22,176 @@ include "basic_2/relocation/lifts.ma".
 (* Basic_2A1: includes: drop_atom drop_pair drop_drop drop_skip
                         drop_refl_atom_O2 drop_drop_lt drop_skip_lt
 *)
-inductive drops (s:bool): trace → relation lenv ≝
-| drops_atom: ∀t. (s = Ⓕ → 𝐈⦃t⦄) → drops s (t) (⋆) (⋆)
-| drops_drop: ∀I,L1,L2,V,t. drops s t L1 L2 → drops s (Ⓕ@t) (L1.ⓑ{I}V) L2
-| drops_skip: ∀I,L1,L2,V1,V2,t.
-              drops s t L1 L2 → ⬆*[t] V2 ≡ V1 →
-              drops s (Ⓣ@t) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
+inductive drops (c:bool): rtmap → relation lenv ≝
+| drops_atom: ∀f. (c = Ⓣ → 𝐈⦃f⦄) → drops c (f) (⋆) (⋆)
+| drops_drop: ∀I,L1,L2,V,f. drops c f L1 L2 → drops c (⫯f) (L1.ⓑ{I}V) L2
+| drops_skip: ∀I,L1,L2,V1,V2,f.
+              drops c f L1 L2 → ⬆*[f] V2 ≡ V1 →
+              drops c (↑f) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2)
 .
 
 interpretation "general slicing (local environment)"
-   'RDropStar s t L1 L2 = (drops s t L1 L2).
+   'RDropStar c f L1 L2 = (drops c f L1 L2).
 
 definition d_liftable1: relation2 lenv term → predicate bool ≝
-                        λR,s. ∀L,K,t. ⬇*[s, t] L ≡ K →
-                        ∀T,U. ⬆*[t] T ≡ U → R K T → R L U.
+                        λR,c. ∀L,K,f. ⬇*[c, f] L ≡ K →
+                        ∀T,U. ⬆*[f] T ≡ U → R K T → R L U.
 
 definition d_liftable2: predicate (lenv → relation term) ≝
-                        λR. ∀K,T1,T2. R K T1 T2 → ∀L,s,t. ⬇*[s, t] L ≡ K →
-                        ∀U1. ⬆*[t] T1 ≡ U1 → 
-                        ∃∃U2. ⬆*[t] T2 ≡ U2 & R L U1 U2.
+                        λR. ∀K,T1,T2. R K T1 T2 → ∀L,c,f. ⬇*[c, f] L ≡ K →
+                        ∀U1. ⬆*[f] T1 ≡ U1 → 
+                        ∃∃U2. ⬆*[f] T2 ≡ U2 & R L U1 U2.
 
 definition d_deliftable2_sn: predicate (lenv → relation term) ≝
-                             λR. ∀L,U1,U2. R L U1 U2 → ∀K,s,t. ⬇*[s, t] L ≡ K →
-                             ∀T1. ⬆*[t] T1 ≡ U1 →
-                             ∃∃T2. ⬆*[t] T2 ≡ U2 & R K T1 T2.
-
-definition dropable_sn: predicate (relation lenv) ≝
-                        λR. ∀L1,K1,s,t. ⬇*[s, t] L1 ≡ K1 → ∀L2. R L1 L2 →
-                        ∃∃K2. R K1 K2 & ⬇*[s, t] L2 ≡ K2.
-(*
-definition dropable_dx: predicate (relation lenv) ≝
-                        λR. ∀L1,L2. R L1 L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 →
-                        ∃∃K1. ⬇[s, 0, m] L1 ≡ K1 & R K1 K2.
-*)
+                             λR. ∀L,U1,U2. R L U1 U2 → ∀K,c,f. ⬇*[c, f] L ≡ K →
+                             ∀T1. ⬆*[f] T1 ≡ U1 →
+                             ∃∃T2. ⬆*[f] T2 ≡ U2 & R K T1 T2.
+
+definition dropable_sn: predicate (rtmap → relation lenv) ≝
+                        λR. ∀L1,K1,c,f. ⬇*[c, f] L1 ≡ K1 → ∀L2,u2. R u2 L1 L2 →
+                        ∀u1. f ⊚ u1 ≡ u2 →
+                        ∃∃K2. R u1 K1 K2 & ⬇*[c, f] L2 ≡ K2.
+
 (* Basic inversion lemmas ***************************************************)
 
-fact drops_inv_atom1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → X = ⋆ →
-                          Y = ⋆ ∧ (s = Ⓕ → 𝐈⦃t⦄).
-#X #Y #s #t * -X -Y -t
+fact drops_inv_atom1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → X = ⋆ →
+                          Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
+#X #Y #c #f * -X -Y -f
 [ /3 width=1 by conj/
-| #I #L1 #L2 #V #t #_ #H destruct
-| #I #L1 #L2 #V1 #V2 #t #_ #_ #H destruct
+| #I #L1 #L2 #V #f #_ #H destruct
+| #I #L1 #L2 #V1 #V2 #f #_ #_ #H destruct
 ]
 qed-.
 
 (* Basic_1: includes: drop_gen_sort *)
 (* Basic_2A1: includes: drop_inv_atom1 *)
-lemma drops_inv_atom1: ∀Y,s,t. ⬇*[s, t] ⋆ ≡ Y → Y = ⋆ ∧ (s = Ⓕ → 𝐈⦃t⦄).
+lemma drops_inv_atom1: ∀Y,c,f. ⬇*[c, f] ⋆ ≡ Y → Y = ⋆ ∧ (c = Ⓣ → 𝐈⦃f⦄).
 /2 width=3 by drops_inv_atom1_aux/ qed-.
 
-fact drops_inv_drop1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → ∀I,K,V,tl. X = K.ⓑ{I}V → t = Ⓕ@tl →
-                          ⬇*[s, tl] K ≡ Y.
-#X #Y #s #t * -X -Y -t
-[ #t #Ht #J #K #W #tl #H destruct
-| #I #L1 #L2 #V #t #HL #J #K #W #tl #H1 #H2 destruct //
-| #I #L1 #L2 #V1 #V2 #t #_ #_ #J #K #W #tl #_ #H2 destruct
+fact drops_inv_drop1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K,V,g. X = K.ⓑ{I}V → f = ⫯g →
+                          ⬇*[c, g] K ≡ Y.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K #W #g #H destruct
+| #I #L1 #L2 #V #f #HL #J #K #W #g #H1 #H2 <(injective_next … H2) -g destruct //
+| #I #L1 #L2 #V1 #V2 #f #_ #_ #J #K #W #g #_ #H2 elim (discr_push_next … H2)
 ]
 qed-.
 
 (* Basic_1: includes: drop_gen_drop *)
 (* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *)
-lemma drops_inv_drop1: ∀I,K,Y,V,s,t. ⬇*[s, Ⓕ@t] K.ⓑ{I}V ≡ Y → ⬇*[s, t] K ≡ Y.
+lemma drops_inv_drop1: ∀I,K,Y,V,c,f. ⬇*[c, ⫯f] K.ⓑ{I}V ≡ Y → ⬇*[c, f] K ≡ Y.
 /2 width=7 by drops_inv_drop1_aux/ qed-.
 
 
-fact drops_inv_skip1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → ∀I,K1,V1,tl. X = K1.ⓑ{I}V1 → t = Ⓣ@tl →
-                          ∃∃K2,V2. ⬇*[s, tl] K1 ≡ K2 & ⬆*[tl] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
-#X #Y #s #t * -X -Y -t
-[ #t #Ht #J #K1 #W1 #tl #H destruct
-| #I #L1 #L2 #V #t #_ #J #K1 #W1 #tl #_ #H2 destruct
-| #I #L1 #L2 #V1 #V2 #t #HL #HV #J #K1 #W1 #tl #H1 #H2 destruct
+fact drops_inv_skip1_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K1,V1,g. X = K1.ⓑ{I}V1 → f = ↑g →
+                          ∃∃K2,V2. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K1 #W1 #g #H destruct
+| #I #L1 #L2 #V #f #_ #J #K1 #W1 #g #_ #H2 elim (discr_next_push … H2)
+| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K1 #W1 #g #H1 #H2 <(injective_push … H2) -g destruct
   /2 width=5 by ex3_2_intro/
 ]
 qed-.
 
 (* Basic_1: includes: drop_gen_skip_l *)
 (* Basic_2A1: includes: drop_inv_skip1 *)
-lemma drops_inv_skip1: ∀I,K1,V1,Y,s,t. ⬇*[s, Ⓣ@t] K1.ⓑ{I}V1 ≡ Y →
-                       ∃∃K2,V2. ⬇*[s, t] K1 ≡ K2 & ⬆*[t] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
+lemma drops_inv_skip1: ∀I,K1,V1,Y,c,f. ⬇*[c, ↑f] K1.ⓑ{I}V1 ≡ Y →
+                       ∃∃K2,V2. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & Y = K2.ⓑ{I}V2.
 /2 width=5 by drops_inv_skip1_aux/ qed-.
 
-fact drops_inv_skip2_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → ∀I,K2,V2,tl. Y = K2.ⓑ{I}V2 → t = Ⓣ@tl →
-                          ∃∃K1,V1. ⬇*[s, tl] K1 ≡ K2 & ⬆*[tl] V2 ≡ V1 & X = K1.ⓑ{I}V1.
-#X #Y #s #t * -X -Y -t
-[ #t #Ht #J #K2 #W2 #tl #H destruct
-| #I #L1 #L2 #V #t #_ #J #K2 #W2 #tl #_ #H2 destruct
-| #I #L1 #L2 #V1 #V2 #t #HL #HV #J #K2 #W2 #tl #H1 #H2 destruct
+fact drops_inv_skip2_aux: ∀X,Y,c,f. ⬇*[c, f] X ≡ Y → ∀I,K2,V2,g. Y = K2.ⓑ{I}V2 → f = ↑g →
+                          ∃∃K1,V1. ⬇*[c, g] K1 ≡ K2 & ⬆*[g] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+#X #Y #c #f * -X -Y -f
+[ #f #Ht #J #K2 #W2 #g #H destruct
+| #I #L1 #L2 #V #f #_ #J #K2 #W2 #g #_ #H2 elim (discr_next_push … H2)
+| #I #L1 #L2 #V1 #V2 #f #HL #HV #J #K2 #W2 #g #H1 #H2 <(injective_push … H2) -g destruct
   /2 width=5 by ex3_2_intro/
 ]
 qed-.
 
 (* Basic_1: includes: drop_gen_skip_r *)
 (* Basic_2A1: includes: drop_inv_skip2 *)
-lemma drops_inv_skip2: ∀I,X,K2,V2,s,t. ⬇*[s, Ⓣ@t] X ≡ K2.ⓑ{I}V2 →
-                       ∃∃K1,V1. ⬇*[s, t] K1 ≡ K2 & ⬆*[t] V2 ≡ V1 & X = K1.ⓑ{I}V1.
+lemma drops_inv_skip2: ∀I,X,K2,V2,c,f. ⬇*[c, ↑f] X ≡ K2.ⓑ{I}V2 →
+                       ∃∃K1,V1. ⬇*[c, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1 & X = K1.ⓑ{I}V1.
 /2 width=5 by drops_inv_skip2_aux/ qed-.
 
 (* Basic properties *********************************************************)
 
+lemma drops_eq_repl_back: ∀L1,L2,c. eq_stream_repl_back … (λf. ⬇*[c, f] L1 ≡ L2).
+#L1 #L2 #c #f1 #H elim H -L1 -L2 -f1
+[ /4 width=3 by drops_atom, isid_eq_repl_back/
+| #I #L1 #L2 #V #f1 #_ #IH #f2 #H elim (next_inv_sn … H) -H
+  /3 width=1 by drops_drop/
+| #I #L1 #L2 #V1 #v2 #f1 #_ #HV #IH #f2 #H elim (push_inv_sn … H) -H
+  /3 width=3 by drops_skip, lifts_eq_repl_back/
+]
+qed-.
+
+lemma drops_eq_repl_fwd: ∀L1,L2,c. eq_stream_repl_fwd … (λf. ⬇*[c, f] L1 ≡ L2).
+#L1 #L2 #c @eq_stream_repl_sym /2 width=3 by drops_eq_repl_back/ (**) (* full auto fails *)
+qed-.
+
+(* Basic_2A1: includes: drop_refl *)
+lemma drops_refl: ∀c,L,f. 𝐈⦃f⦄ → ⬇*[c, f] L ≡ L.
+#c #L elim L -L /2 width=1 by drops_atom/
+#L #I #V #IHL #f #Hf elim (isid_inv_gen … Hf) -Hf
+/3 width=1 by drops_skip, lifts_refl/
+qed.
+
 (* Basic_2A1: includes: drop_FT *)
-lemma drops_FT: ∀L1,L2,t. ⬇*[Ⓕ, t] L1 ≡ L2 → ⬇*[Ⓣ, t] L1 ≡ L2.
-#L1 #L2 #t #H elim H -L1 -L2 -t
+lemma drops_TF: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
+#L1 #L2 #f #H elim H -L1 -L2 -f
 /3 width=1 by drops_atom, drops_drop, drops_skip/
 qed.
 
 (* Basic_2A1: includes: drop_gen *)
-lemma drops_gen: ∀L1,L2,s,t. ⬇*[Ⓕ, t] L1 ≡ L2 → ⬇*[s, t] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by drops_FT/
+lemma drops_gen: ∀L1,L2,c,f. ⬇*[Ⓣ, f] L1 ≡ L2 → ⬇*[c, f] L1 ≡ L2.
+#L1 #L2 * /2 width=1 by drops_TF/
 qed-.
 
 (* Basic_2A1: includes: drop_T *)
-lemma drops_T: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → ⬇*[Ⓣ, t] L1 ≡ L2.
-#L1 #L2 * /2 width=1 by drops_FT/
+lemma drops_F: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → ⬇*[Ⓕ, f] L1 ≡ L2.
+#L1 #L2 * /2 width=1 by drops_TF/
 qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-fact drops_fwd_drop2_aux: ∀X,Y,s,t2. ⬇*[s, t2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
-                          ∃∃t1,t. 𝐈⦃t1⦄ & t2 ⊚ Ⓕ@t1 ≡ t & ⬇*[s, t] X ≡ K.
-#X #Y #s #t2 #H elim H -X -Y -t2
-[ #t2 #Ht2 #J #K #W #H destruct
-| #I #L1 #L2 #V #t2 #_ #IHL #J #K #W #H elim (IHL … H) -IHL
-  /3 width=5 by after_false, ex3_2_intro, drops_drop/
-| #I #L1 #L2 #V1 #V2 #t2 #HL #_ #_ #J #K #W #H destruct
-  elim (isid_after_dx t2) /3 width=5 by after_true, ex3_2_intro, drops_drop/
+fact drops_fwd_drop2_aux: ∀X,Y,c,f2. ⬇*[c, f2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
+                          ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[c, f] X ≡ K.
+#X #Y #c #f2 #H elim H -X -Y -f2
+[ #f2 #Ht2 #J #K #W #H destruct
+| #I #L1 #L2 #V #f2 #_ #IHL #J #K #W #H elim (IHL … H) -IHL
+  /3 width=7 by after_next, ex3_2_intro, drops_drop/
+| #I #L1 #L2 #V1 #V2 #f2 #HL #_ #_ #J #K #W #H destruct
+  lapply (isid_after_dx 𝐈𝐝 f2 ?) /3 width=9 by after_push, ex3_2_intro, drops_drop/
 ]
 qed-.
 
 (* Basic_1: includes: drop_S *)
 (* Basic_2A1: includes: drop_fwd_drop2 *)
-lemma drops_fwd_drop2: ∀I,X,K,V,s,t2. ⬇*[s, t2] X ≡ K.ⓑ{I}V →
-                       ∃∃t1,t. 𝐈⦃t1⦄ & t2 ⊚ Ⓕ@t1 ≡ t & ⬇*[s, t] X ≡ K.
+lemma drops_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
+                       ∃∃f1,f. 𝐈⦃f1⦄ & f2 ⊚ ⫯f1 ≡ f & ⬇*[c, f] X ≡ K.
 /2 width=5 by drops_fwd_drop2_aux/ qed-.
 
-fact drops_after_fwd_drop2_aux: ∀X,Y,s,t2. ⬇*[s, t2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V →
-                                ∀t1,t. 𝐈⦃t1⦄ → t2 ⊚ Ⓕ@t1 ≡ t → ⬇*[s, t] X ≡ K.
-#X #Y #s #t2 #H elim H -X -Y -t2
-[ #t2 #Ht2 #J #K #W #H destruct
-| #I #L1 #L2 #V #t2 #_ #IHL #J #K #W #H #t1 #t #Ht1 #Ht elim (after_inv_false1 … Ht) -Ht
-  /3 width=3 by drops_drop/
-| #I #L1 #L2 #V1 #V2 #t2 #HL #_ #_ #J #K #W #H #t1 #t #Ht1 #Ht elim (after_inv_true1 … Ht) -Ht
-  #u1 #u #b #H1 #H2 #Hu destruct >(after_isid_inv_dx … Hu) -Hu /2 width=1 by drops_drop/
-]
+lemma drops_after_fwd_drop2: ∀I,X,K,V,c,f2. ⬇*[c, f2] X ≡ K.ⓑ{I}V →
+                             ∀f1,f. 𝐈⦃f1⦄ → f2 ⊚ ⫯f1 ≡ f → ⬇*[c, f] X ≡ K.
+#I #X #K #V #c #f2 #H #f1 #f #Hf1 #Hf elim (drops_fwd_drop2 … H) -H
+#g1 #g #Hg1 #Hg #HK lapply (after_mono … Hg Hf ??) -Hg -Hf
+/3 width=3 by drops_eq_repl_back, isid_inv_eq_repl, next_eq_repl/
 qed-.
 
-lemma drops_after_fwd_drop2: ∀I,X,K,V,s,t2. ⬇*[s, t2] X ≡ K.ⓑ{I}V →
-                             ∀t1,t. 𝐈⦃t1⦄ → t2 ⊚ Ⓕ@t1 ≡ t → ⬇*[s, t] X ≡ K.
-/2 width=9 by drops_after_fwd_drop2_aux/ qed-.
-
 (* Basic_1: includes: drop_gen_refl *)
 (* Basic_2A1: includes: drop_inv_O2 *)
-lemma drops_fwd_isid: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → 𝐈⦃t⦄ → L1 = L2.
-#L1 #L2 #s #t #H elim H -L1 -L2 -t //
-[ #I #L1 #L2 #V #t #_ #_ #H elim (isid_inv_false … H)
-| /5 width=3 by isid_inv_true, lifts_fwd_isid, eq_f3, sym_eq/
+lemma drops_fwd_isid: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → 𝐈⦃f⦄ → L1 = L2.
+#L1 #L2 #c #f #H elim H -L1 -L2 -f //
+[ #I #L1 #L2 #V #f #_ #_ #H elim (isid_inv_next … H)
+| /5 width=3 by isid_inv_push, lifts_fwd_isid, eq_f3, sym_eq/
 ]
 qed-.
 
-(* Basic_2A1: removed theorems 13:
+(* Basic_2A1: removed theorems 14:
               drops_inv_nil drops_inv_cons d1_liftable_liftables
+              drop_refl_atom_O2
               drop_inv_O1_pair1 drop_inv_pair1 drop_inv_O1_pair2
               drop_inv_Y1 drop_Y1 drop_O_Y drop_fwd_Y2
               drop_fwd_length_minus2 drop_fwd_length_minus4
index 34c446f649b8d5863693a4194d2839be57dafefa..56b39051f670f6008ef5778b58ad0ef6e9feb31c 100644 (file)
@@ -20,18 +20,18 @@ include "basic_2/relocation/drops.ma".
 (* Main properties **********************************************************)
 
 (* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *)
-theorem drops_conf: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L →
-                    ∀L2,s2,t. ⬇*[s2, t] L1 ≡ L2 →
-                    ∀t2. t1 ⊚ t2 ≡ t → ⬇*[s2, t2] L ≡ L2.
-#L1 #L #s1 #t1 #H elim H -L1 -L -t1
-[ #t1 #_ #L2 #s2 #t #H #t2 #Ht12 elim (drops_inv_atom1 … H) -s1 -H
-  #H #Ht destruct @drops_atom
-  #H elim (after_inv_isid3 … Ht12) -Ht12 /2 width=1 by/
-| #I #K1 #K #V1 #t1 #_ #IH #L2 #s2 #t #H12 #t2 #Ht elim (after_inv_false1 … Ht) -Ht
-  #u #H #Hu destruct /3 width=3 by drops_inv_drop1/
-| #I #K1 #K #V1 #V #t1 #_ #HV1 #IH #L2 #s2 #t #H #t2 #Ht elim (after_inv_true1 … Ht) -Ht
-  #u2 #u * #H1 #H2 #Hu destruct
-  [ elim (drops_inv_skip1 … H) -H /3 width=6 by drops_skip, lifts_div/
+theorem drops_conf: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
+                    ∀L2,c2,f. ⬇*[c2, f] L1 ≡ L2 →
+                    ∀f2. f1 ⊚ f2 ≡ f → ⬇*[c2, f2] L ≡ L2.
+#L1 #L #c1 #f1 #H elim H -L1 -L -f1
+[ #f1 #_ #L2 #c2 #f #HL2 #f2 #Hf12 elim (drops_inv_atom1 … HL2) -c1 -HL2
+  #H #Hf destruct @drops_atom
+  #H elim (after_inv_isid3 … Hf12) -Hf12 /2 width=1 by/
+| #I #K1 #K #V1 #f1 #_ #IH #L2 #c2 #f #HL2 #f2 #Hf elim (after_inv_Sxx … Hf) -Hf
+  #g #Hg #H destruct /3 width=3 by drops_inv_drop1/
+| #I #K1 #K #V1 #V #f1 #_ #HV1 #IH #L2 #c2 #f #HL2 #f2 #Hf elim (after_inv_Oxx … Hf) -Hf *
+  #g2 #g #Hf #H1 #H2 destruct
+  [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_div/
   | /4 width=3 by drops_inv_drop1, drops_drop/
   ]
 ]
@@ -41,18 +41,19 @@ qed-.
 (* Basic_2A1: includes: drop_trans_ge drop_trans_le drop_trans_ge_comm 
                         drops_drop_trans
 *)
-theorem drops_trans: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L →
-                     ∀L2,s2,t2. ⬇*[s2, t2] L ≡ L2 →
-                     ∀t. t1 ⊚ t2 ≡ t → ⬇*[s1∨s2, t] L1 ≡ L2.
-#L1 #L #s1 #t1 #H elim H -L1 -L -t1
-[ #t1 #Ht1 #L2 #s2 #t2 #H #t #Ht elim (drops_inv_atom1 … H) -H
-  #H #Ht2 destruct @drops_atom #H elim (orb_false_r … H) -H
-  #H1 #H2 >(after_isid_inv_sn … Ht) -Ht /2 width=1 by/
-| #I #K1 #K #V1 #t1 #_ #IH #L #s2 #t2 #HKL #t #Ht elim (after_inv_false1 … Ht) -Ht
+theorem drops_trans: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
+                     ∀L2,c2,f2. ⬇*[c2, f2] L ≡ L2 →
+                     ∀f. f1 ⊚ f2 ≡ f → ⬇*[c1∧c2, f] L1 ≡ L2.
+#L1 #L #c1 #f1 #H elim H -L1 -L -f1
+[ #f1 #Hf1 #L2 #c2 #f2 #HL2 #f #Hf elim (drops_inv_atom1 … HL2) -HL2
+  #H #Hf2 destruct @drops_atom #H elim (andb_inv_true_dx … H) -H
+  #H1 #H2 lapply (after_isid_inv_sn … Hf ?) -Hf
+  /3 width=3 by isid_eq_repl_back/
+| #I #K1 #K #V1 #f1 #_ #IH #L2 #c2 #f2 #HL2 #f #Hf elim (after_inv_Sxx … Hf) -Hf
   /3 width=3 by drops_drop/
-| #I #K1 #K #V1 #V #t1 #_ #HV1 #IH #L #s2 #t2 #H #t #Ht elim (after_inv_true1 … Ht) -Ht
-  #u2 #u * #H1 #H2 #Hu destruct
-  [ elim (drops_inv_skip1 … H) -H /3 width=6 by drops_skip, lifts_trans/
+| #I #K1 #K #V1 #V #f1 #_ #HV1 #IH #L2 #c2 #f2 #HL2 #f #Hf elim (after_inv_Oxx … Hf) -Hf *
+  #g2 #g #Hg #H1 #H2 destruct
+  [ elim (drops_inv_skip1 … HL2) -HL2 /3 width=6 by drops_skip, lifts_trans/
   | /4 width=3 by drops_inv_drop1, drops_drop/
   ]
 ]
@@ -61,29 +62,29 @@ qed-.
 (* Advanced properties ******************************************************)
 
 (* Basic_2A1: includes: drop_mono *)
-lemma drops_mono: ∀L,L1,s1,t. ⬇*[s1, t] L ≡ L1 →
-                  ∀L2,s2. ⬇*[s2, t] L ≡ L2 → L1 = L2.
-#L #L1 #s1 #t elim (isid_after_dx t)
+lemma drops_mono: ∀L,L1,c1,f. ⬇*[c1, f] L ≡ L1 →
+                  ∀L2,c2. ⬇*[c2, f] L ≡ L2 → L1 = L2.
+#L #L1 #c1 #f lapply (isid_after_dx 𝐈𝐝 f ?)
 /3 width=8 by drops_conf, drops_fwd_isid/
 qed-.
 
 (* Basic_2A1: includes: drop_conf_lt *)
-lemma drops_conf_skip1: ∀L,L2,s2,t. ⬇*[s2, t] L ≡ L2 →
-                        ∀I,K1,V1,s1,t1. ⬇*[s1, t1] L ≡ K1.ⓑ{I}V1 →
-                        ∀t2. t1 ⊚ Ⓣ@t2 ≡ t →
+lemma drops_conf_skip1: ∀L,L2,c2,f. ⬇*[c2, f] L ≡ L2 →
+                        ∀I,K1,V1,c1,f1. ⬇*[c1, f1] L ≡ K1.ⓑ{I}V1 →
+                        ∀f2. f1 ⊚ ↑f2 ≡ f →
                         ∃∃K2,V2. L2 = K2.ⓑ{I}V2 &
-                                 ⬇*[s2, t2] K1 ≡ K2 & ⬆*[t2] V2 ≡ V1.
-#L #L2 #s2 #t #H2 #I #K1 #V1 #s1 #t1 #H1 #t2 #Ht lapply (drops_conf … H1 … H2 … Ht) -L -Ht
+                                 ⬇*[c2, f2] K1 ≡ K2 & ⬆*[f2] V2 ≡ V1.
+#L #L2 #c2 #f #H2 #I #K1 #V1 #c1 #f1 #H1 #f2 #Hf lapply (drops_conf … H1 … H2 … Hf) -L -Hf
 #H elim (drops_inv_skip1 … H) -H /2 width=5 by ex3_2_intro/
 qed-.
 
 (* Basic_2A1: includes: drop_trans_lt *)
-lemma drops_trans_skip2: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L →
-                         ∀I,K2,V2,s2,t2. ⬇*[s2, t2] L ≡ K2.ⓑ{I}V2 →
-                         ∀t. t1 ⊚ t2 ≡ Ⓣ@t →
+lemma drops_trans_skip2: ∀L1,L,c1,f1. ⬇*[c1, f1] L1 ≡ L →
+                         ∀I,K2,V2,c2,f2. ⬇*[c2, f2] L ≡ K2.ⓑ{I}V2 →
+                         ∀f. f1 ⊚ f2 ≡ ↑f →
                          ∃∃K1,V1. L1 = K1.ⓑ{I}V1 &
-                                  ⬇*[s1∨s2, t] K1 ≡ K2 & ⬆*[t] V2 ≡ V1.
-#L1 #L #s1 #t1 #H1 #I #K2 #V2 #s2 #t2 #H2 #t #Ht
-lapply (drops_trans … H1 … H2 … Ht) -L -Ht
+                                  ⬇*[c1∧c2, f] K1 ≡ K2 & ⬆*[f] V2 ≡ V1.
+#L1 #L #c1 #f1 #H1 #I #K2 #V2 #c2 #f2 #H2 #f #Hf
+lapply (drops_trans … H1 … H2 … Hf) -L -Hf
 #H elim (drops_inv_skip2 … H) -H /2 width=5 by ex3_2_intro/
 qed-.
index 5851e809e219742eccb201d9cd3a2772071bc08f..338ef22d5a5a0514db8a9f634f801ecfcea39317 100644 (file)
@@ -22,9 +22,9 @@ include "basic_2/relocation/drops.ma".
 (* Basic_2A1: was: d_liftable_LTC *)
 lemma d2_liftable_LTC: ∀R. d_liftable2 R → d_liftable2 (LTC … R).
 #R #HR #K #T1 #T2 #H elim H -T2
-[ #T2 #HT12 #L #s #t #HLK #U1 #HTU1
+[ #T2 #HT12 #L #c #f #HLK #U1 #HTU1
   elim (HR … HT12 … HLK … HTU1) /3 width=3 by inj, ex2_intro/
-| #T #T2 #_ #HT2 #IHT1 #L #s #t #HLK #U1 #HTU1
+| #T #T2 #_ #HT2 #IHT1 #L #c #f #HLK #U1 #HTU1
   elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1
   elim (HR … HT2 … HLK … HTU) -HR -K -T /3 width=5 by step, ex2_intro/
 ]
@@ -33,38 +33,29 @@ qed-.
 (* Basic_2A1: was: d_deliftable_sn_LTC *)
 lemma d2_deliftable_sn_LTC: ∀R. d_deliftable2_sn R → d_deliftable2_sn (LTC … R).
 #R #HR #L #U1 #U2 #H elim H -U2
-[ #U2 #HU12 #K #s #t #HLK #T1 #HTU1
+[ #U2 #HU12 #K #c #f #HLK #T1 #HTU1
   elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/
-| #U #U2 #_ #HU2 #IHU1 #K #s #t #HLK #T1 #HTU1
+| #U #U2 #_ #HU2 #IHU1 #K #c #f #HLK #T1 #HTU1
   elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
   elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/
 ]
 qed-.
 
-lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (TC … R).
-#R #HR #L1 #K1 #s #t #HLK1 #L2 #H elim H -L2
-[ #L2 #HL12 elim (HR … HLK1 … HL12) -HR -L1
+lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (LTC … R).
+#R #HR #L1 #K1 #c #f #HLK1 #L2 #u2 #H elim H -L2
+[ #L2 #HL12 #u1 #H elim (HR … HLK1 … HL12 … H) -HR -L1 -u2
   /3 width=3 by inj, ex2_intro/
-| #L #L2 #_ #HL2 * #K #HK1 #HLK elim (HR … HLK … HL2) -HR -L
+| #L #L2 #_ #HL2 #IH #u1 #H elim (IH … H) -IH
+  #K #HK1 #HLK elim (HR … HLK … HL2 … H) -HR -L -u2
   /3 width=3 by step, ex2_intro/
 ]
 qed-.
-(*
-lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R).
-#R #HR #L1 #L2 #H elim H -L2
-[ #L2 #HL12 #K2 #s #m #HLK2 elim (HR … HL12 … HLK2) -HR -L2
-  /3 width=3 by inj, ex2_intro/
-| #L #L2 #_ #HL2 #IHL1 #K2 #s #m #HLK2 elim (HR … HL2 … HLK2) -HR -L2
-  #K #HLK #HK2 elim (IHL1 … HLK) -L
-  /3 width=5 by step, ex2_intro/
-]
-qed-.
-*)
+
 (* Basic_2A1: was: d_liftable_llstar *)
 lemma d2_liftable_llstar: ∀R. d_liftable2 R → ∀d. d_liftable2 (llstar … R d).
 #R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
-[ #L #s #t #_ #U1 #HTU1 -HR -K -s /2 width=3 by ex2_intro/
-| #d #T #T2 #_ #HT2 #IHT1 #L #s #t #HLK #U1 #HTU1
+[ #L #c #f #_ #U1 #HTU1 -HR -K -c /2 width=3 by ex2_intro/
+| #d #T #T2 #_ #HT2 #IHT1 #L #c #f #HLK #U1 #HTU1
   elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1
   elim (HR … HT2 … HLK … HTU) -T /3 width=5 by lstar_dx, ex2_intro/
 ]
@@ -75,7 +66,7 @@ lemma d2_deliftable_sn_llstar: ∀R. d_deliftable2_sn R →
                                ∀d. d_deliftable2_sn (llstar … R d).
 #R #HR #d #L #U1 #U2 #H @(lstar_ind_r … d U2 H) -d -U2
 [ /2 width=3 by lstar_O, ex2_intro/
-| #d #U #U2 #_ #HU2 #IHU1 #K #s #t #HLK #T1 #HTU1
+| #d #U #U2 #_ #HU2 #IHU1 #K #c #f #HLK #T1 #HTU1
   elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1
   elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/
 ]
index 63bc71da6d090c3ad6dec5b4802891b400b5da5a..9cf6e7c1fd76399eac10a7bc48d9b0a663e5e21b 100644 (file)
@@ -18,14 +18,14 @@ include "basic_2/relocation/drops.ma".
 (* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************)
 
 definition d_liftable1_all: relation2 lenv term → predicate bool ≝
-                            λR,s. ∀L,K,t. ⬇*[s, t] L ≡ K →
-                            ∀Ts,Us. ⬆*[t] Ts ≡ Us →
+                            λR,c. ∀L,K,f. ⬇*[c, f] L ≡ K →
+                            ∀Ts,Us. ⬆*[f] Ts ≡ Us →
                             all … (R K) Ts → all … (R L) Us.
 
 (* Properties on general relocation for term vectors ************************)
 
 (* Basic_2A1: was: d1_liftables_liftables_all *)
-lemma d1_liftable_liftable_all: ∀R,s. d_liftable1 R s → d_liftable1_all R s.
-#R #s #HR #L #K #t #HLK #Ts #Us #H elim H -Ts -Us normalize //
+lemma d1_liftable_liftable_all: ∀R,c. d_liftable1 R c → d_liftable1_all R c.
+#R #c #HR #L #K #f #HLK #Ts #Us #H elim H -Ts -Us normalize //
 #Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/
 qed.
index 0db32e2a21950cf7b110ed1d7667db24f4f950fd..0116254957011dd23e186b98f7f4ca96ad8723e6 100644 (file)
@@ -21,30 +21,30 @@ include "basic_2/relocation/drops.ma".
 (* Forward lemmas on weight for local environments **************************)
 
 (* Basic_2A1: includes: drop_fwd_lw *)
-lemma drops_fwd_lw: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
-#L1 #L2 #s #t #H elim H -L1 -L2 -t //
+lemma drops_fwd_lw: ∀L1,L2,c,f. ⬇*[c, f] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
+#L1 #L2 #c #f #H elim H -L1 -L2 -f //
 [ /2 width=3 by transitive_le/
-| #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 normalize
+| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 normalize
   >(lifts_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/
 ]
 qed-.
 
 (* Basic_2A1: includes: drop_fwd_lw_lt *)
 (* Note: 𝐈⦃t⦄ → ⊥ is ∥l∥ < |L| *)
-lemma drops_fwd_lw_lt: ∀L1,L2,t. ⬇*[Ⓕ, t] L1 ≡ L2 →
-                       (𝐈⦃t⦄ → ⊥) → ♯{L2} < ♯{L1}.
-#L1 #L2 #t #H elim H -L1 -L2 -t
-[ #t #Ht #Hnt elim Hnt -Hnt /2 width=1 by/
+lemma drops_fwd_lw_lt: ∀L1,L2,f. ⬇*[Ⓣ, f] L1 ≡ L2 →
+                       (𝐈⦃f⦄ → ⊥) → ♯{L2} < ♯{L1}.
+#L1 #L2 #f #H elim H -L1 -L2 -f
+[ #f #Hf #Hnf elim Hnf -Hnf /2 width=1 by/
 | /3 width=3 by drops_fwd_lw, le_to_lt_to_lt/
-| #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
-  >(lifts_fwd_tw … HV21) -V2 /4 width=1 by monotonic_lt_plus_l/
+| #I #L1 #L2 #V1 #V2 #f #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I
+  >(lifts_fwd_tw … HV21) -V2 /5 width=1 by isid_push, monotonic_lt_plus_l/
 ]
 qed-.
 
 (* Forward lemmas on restricted weight for closures *************************)
 
 (* Basic_2A1: includes: drop_fwd_rfw *)
-lemma drops_pair2_fwd_rfw: ∀I,L,K,V,s,t. ⬇*[s, t] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
-#I #L #K #V #s #t #HLK lapply (drops_fwd_lw … HLK) -HLK
+lemma drops_pair2_fwd_rfw: ∀I,L,K,V,c,f. ⬇*[c, f] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}.
+#I #L #K #V #c #f #HLK lapply (drops_fwd_lw … HLK) -HLK
 normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/
 qed-.
index 5a80b945f98a10c075234a47d29c80da19be7dc0..d4f38b0baead585e99570b7caff1e34cb80ce0d3 100644 (file)
@@ -1581,7 +1581,7 @@ let predefined_classes = [
  ["w"; "ω"; "𝕨"; "𝐰"; "𝛚"; "ⓦ"; ] ;
  ["W"; "Ω"; "𝕎"; "𝐖"; "𝛀"; "Ⓦ"; ] ;
  ["x"; "ξ"; "χ"; "ϰ"; "𝕩"; "𝐱"; "𝛏"; "𝛘"; "𝛞"; "ⓧ"; ] ;
- ["X"; "Ξ"; "𝕏";"𝐗"; "𝚵"; "Ⓧ"; ] ;
+ ["X"; "Ξ"; "𝕏";"𝐗"; "𝚵"; "Ⓧ"; "⦻"; ] ;
  ["y"; "υ"; "𝕪"; "𝐲"; "ⓨ"; ] ;
  ["Y"; "ϒ"; "𝕐"; "𝐘"; "𝚼"; "Ⓨ"; ] ;
  ["z"; "ζ"; "𝕫"; "𝐳"; "𝛇"; "ⓩ"; ] ;