]> matita.cs.unibo.it Git - helm.git/commitdiff
- confluence of parallel substitution (tps) closed! (a bug in an
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 23 Aug 2011 17:07:23 +0000 (17:07 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Tue, 23 Aug 2011 17:07:23 +0000 (17:07 +0000)
inversion lemma was in the way)
- some refactoring

matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma [deleted file]
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_tps.ma
matita/matita/contribs/lambda-delta/Ground-2/xoa.conf.xml
matita/matita/contribs/lambda-delta/Ground-2/xoa.ma
matita/matita/contribs/lambda-delta/Ground-2/xoa_notation.ma

index bc82e4e6d2c870a499c1dcb6ed6bba96e5f36d1a..fc0851097de3709970bb96a922bcdcbffd8e06d0 100644 (file)
@@ -86,27 +86,52 @@ lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
 lapply (tps_weak_top … HT12) //
 qed.
 
+lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i → i ≤ d + e →
+                    ∃∃T. L ⊢ T1 [d, i - d] ≫ T & L ⊢ T [i, d + e - i] ≫ T2.
+#L #T1 #T2 #d #e #H elim H -L T1 T2 d e
+[ /2/
+| /2/
+| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde
+  elim (lt_or_ge i j)
+  [ -Hide Hjde;
+    >(plus_minus_m_m_comm j d) in ⊢ (% → ?) // -Hdj /3/
+  | -Hdi Hdj; #Hid
+    generalize in match Hide -Hide (**) (* rewriting in the premises, rewrites in the goal too *)
+    >(plus_minus_m_m_comm … Hjde) in ⊢ (% → ?) -Hjde /4/
+  ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+  elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
+  elim (IHT12 (i + 1) ? ?) -IHT12 [2: /2 by arith4/ |3: /2/ ] (* just /2/ is too slow *)
+  -Hdi Hide >arith_c1 >arith_c1x #T #HT1 #HT2
+  @ex2_1_intro [2,3: @tps_bind | skip ] (**) (* explicit constructors *)
+  [3: @HV1 |4: @HT1 |5: // |1,2: skip | /3 width=5/ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+  elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 //
+  -Hdi Hide /3 width=5/
+]
+qed.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma tps_inv_lref1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. T1 = #i →
-                         T2 = #i ∨ 
-                         ∃∃K,V,i. d ≤ i & i < d + e &
-                                  ↓[O, i] L ≡ K. 𝕓{Abbr} V &
-                                  ↑[O, i + 1] V ≡ T2.
+                         T2 = #i ∨
+                         ∃∃K,V. d ≤ i & i < d + e &
+                                ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                                ↑[O, i + 1] V ≡ T2.
 #L #T1 #T2 #d #e * -L T1 T2 d e
 [ #L #k #d #e #i #H destruct
 | /2/
-| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #j #H destruct -i /3 width=7/
+| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #j #H destruct -i /3/
 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
 ]
 qed.
 
 lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ≫ T2 →
-                     T2 = #i ∨ 
-                     ∃∃K,V,i. d ≤ i & i < d + e &
-                              ↓[O, i] L ≡ K. 𝕓{Abbr} V &
-                              ↑[O, i + 1] V ≡ T2.
+                     T2 = #i ∨
+                     ∃∃K,V. d ≤ i & i < d + e &
+                            ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                            ↑[O, i + 1] V ≡ T2.
 /2/ qed.
 
 lemma tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
index e5ad25ff13ebf98ea6fccfb4df3a3560389e39fc..d15b59535d33e6d678e22afb9f00667b397684dc 100644 (file)
@@ -170,3 +170,14 @@ qed.
                                         (le d i) -> (lt i (plus d h)) ->
                                        (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)).
 *)
+
+lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
+                        ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
+                        d ≤ dt → dt ≤ d + e → d + e ≤ dt + et →
+                        ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2.
+#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
+elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2
+lapply (tps_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #HU1
+lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -U1;
+elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 HLK HTU1 // <minus_plus_m_m /2/
+qed.
diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma
deleted file mode 100644 (file)
index 6bb476f..0000000
+++ /dev/null
@@ -1,55 +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/substitution/tps_lift.ma".
-
-(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
-
-(* Split properties *********************************************************)
-
-lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i → i ≤ d + e →
-                    ∃∃T. L ⊢ T1 [d, i - d] ≫ T & L ⊢ T [i, d + e - i] ≫ T2.
-#L #T1 #T2 #d #e #H elim H -L T1 T2 d e
-[ /2/
-| /2/
-| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde
-  elim (lt_or_ge i j)
-  [ -Hide Hjde;
-    >(plus_minus_m_m_comm j d) in ⊢ (% → ?) // -Hdj /3/
-  | -Hdi Hdj; #Hid
-    generalize in match Hide -Hide (**) (* rewriting in the premises, rewrites in the goal too *)
-    >(plus_minus_m_m_comm … Hjde) in ⊢ (% → ?) -Hjde /4/
-  ]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
-  elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
-  elim (IHT12 (i + 1) ? ?) -IHT12 [2: /2 by arith4/ |3: /2/ ] (* just /2/ is too slow *)
-  -Hdi Hide >arith_c1 >arith_c1x #T #HT1 #HT2
-  @ex2_1_intro [2,3: @tps_bind | skip ] (**) (* explicit constructors *)
-  [3: @HV1 |4: @HT1 |5: // |1,2: skip | /3 width=5/ ]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
-  elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 //
-  -Hdi Hide /3 width=5/
-]
-qed.
-
-lemma tps_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
-                        ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
-                        d ≤ dt → dt ≤ d + e → d + e ≤ dt + et →
-                        ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2.
-#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
-elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2
-lapply (tps_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #HU1
-lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -U1;
-elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 HLK HTU1 // <minus_plus_m_m /2/
-qed.
index 1e3c3629a189e6ba42f96c5f3e130de5e235f5e1..83316ac30a429fa907561ba86804f3fc0409a969 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic-2/substitution/tps_split.ma".
+include "Basic-2/substitution/tps_lift.ma".
 
-(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+(* PARALLEL SUBSTITUTION ON TERMS *******************************************)
 
 (* Main properties **********************************************************)
-(*
-theorem tps_trans: ∀L,T1,T,d,e. L ⊢ T1 [d, e] ≫ T → ∀T2. L ⊢ T [d, e] ≫ T2 →
-                   L ⊢ T1 [d, e] ≫ T2.
-#L #T1 #T #d #e #H elim H -L T1 T d e
-[ //
-| //
-| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #_ #HV12 #IHV12 #T2 #HVT2
-  lapply (drop_fwd_drop2 … HLK) #H
-  elim (tps_inv_lift1_up … HVT2 … H … HV12 ? ? ?) -HVT2 H HV12 // normalize [2: /2/ ] #V #HV1 #HVT2
-  @tps_subst [4,5,6,8: // |1,2,3: skip | /2/ ]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
-  elim (tps_inv_bind1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X;
-  @tps_bind /2/ @IHT12 @(tps_leq_repl … HT2) /2/
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
-  elim (tps_inv_flat1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X /3/
-]
-qed.
-*)
-
-axiom tps_conf_subst_subst_lt: ∀L,K1,V1,T1,i1,d,e,T2,K2,V2,i2.
-   ↓[O, i1] L ≡ K1. 𝕓{Abbr} V1 → ↓[O, i2] L≡ K2. 𝕓{Abbr} V2 →
-   ↑[O, i1 + 1] V1 ≡ T1 → ↑[O, i2 + 1] V2 ≡ T2 → 
-   d ≤ i1 → i1 < d + e → d ≤ i2 → i2 < d + e → i1 < i2 →
-   ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T.  
-(*
-#L #K1 #V1 #T1 #i1 #d #e #T2 #K2 #V2 #i2
-#HLK1 #HLK2 #HVW1 #HVW2 #HWT1 #HWT2 #Hdi1 #Hi1de #Hdi2 #Hi2de #Hi12
-*)
 
 theorem tps_conf: ∀L,T0,T1,d,e. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d, e] ≫ T2 →
                   ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T.
@@ -54,31 +26,48 @@ theorem tps_conf: ∀L,T0,T1,d,e. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d,
 | #L #K1 #V1 #T1 #i1 #d #e #Hdi1 #Hi1de #HLK1 #HVT1 #T2 #H
   elim (tps_inv_lref1 … H) -H
   [ #HX destruct -T2 /4/
-  | * #K2 #V2 #i2 #Hdi2 #Hi2de #HLK2 #HVT2
-    elim (lt_or_eq_or_gt i1 i2) #Hi12
-    [ @tps_conf_subst_subst_lt /width=15/
-    | -Hdi2 Hi2de; destruct -i2;
-      lapply (drop_mono … HLK1 … HLK2) -HLK1 #H destruct -V1 K1
-      >(lift_mono … HVT1 … HVT2) -HVT1 /2/
-    | @ex2_1_comm @tps_conf_subst_subst_lt /width=15/
-    ]
+  | * #K2 #V2 #_ #_ #HLK2 #HVT2
+    lapply (drop_mono … HLK1 … HLK2) -HLK1 #H destruct -V1 K1
+    >(lift_mono … HVT1 … HVT2) -HVT1 /2/
   ]
 | #L #I #V0 #V1 #T0 #T1 #d #e #_ #_ #IHV01 #IHT01 #X #HX
-  elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; 
+  elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X;
   elim (IHV01 … HV02) -IHV01 HV02 #V #HV1 #HV2
-  elim (IHT01 … HT02) -IHT01 HT02 #T #HT1 #HT2 
+  elim (IHT01 … HT02) -IHT01 HT02 #T #HT1 #HT2
   @ex2_1_intro
   [2: @tps_bind [4: @(tps_leq_repl … HT1) /2/ |2: skip ]
   |1: skip
   |3: @tps_bind [2: @(tps_leq_repl … HT2) /2/ ]
   ] // (**) (* /5/ is too slow *)
 | #L #I #V0 #V1 #T0 #T1 #d #e #_ #_ #IHV01 #IHT01 #X #HX
-  elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X; 
+  elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X;
   elim (IHV01 … HV02) -IHV01 HV02;
   elim (IHT01 … HT02) -IHT01 HT02 /3 width=5/
 ]
 qed.
 
+theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ≫ T0 →
+                        ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫ T2 → d2 + e2 ≤ d1 →
+                        ∃∃T. L ⊢ T1 [d2, e2] ≫ T & L ⊢ T [d1, e1] ≫ T2.
+#L #T1 #T0 #d1 #e1 #H elim H -L T1 T0 d1 e1
+[ /2/
+| /2/
+| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1
+  lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1
+  lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2; normalize /2/ -Hde2i1 #HWT2
+  <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /4/
+| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
+  elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X;
+  lapply (tps_leq_repl … HT02 (L. 𝕓{I} V1) ?) -HT02 /2/ #HT02
+  elim (IHV10 … HV02 ?) -IHV10 HV02 // #V
+  elim (IHT10 … HT02 ?) -IHT10 HT02 [2: /2/ ] #T #HT1 #HT2
+  lapply (tps_leq_repl … HT2 (L. 𝕓{I} V) ?) -HT2 /3 width=6/
+| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1
+  elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct -X;
+  elim (IHV10 … HV02 ?) -IHV10 HV02 //
+  elim (IHT10 … HT02 ?) -IHT10 HT02 // /3 width=6/
+]
+qed.
 (*
       Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) ->
                              (u1,u:?; i:?) (subst0 i u u1 u2) ->
index f0ba4ec8175d4547c52afa845793cc913331f8fa..e7f9feef45d11d57f922eae013e7959b3a75f5d3 100644 (file)
@@ -19,6 +19,7 @@
     <key name="ex">3 1</key>
     <key name="ex">3 2</key>
     <key name="ex">3 3</key>
+    <key name="ex">4 2</key>
     <key name="ex">4 3</key>
     <key name="ex">4 4</key>
     <key name="ex">5 3</key>
index 68a0776dcdecd17a28c508befb92c5a5303452d0..a9deb0fa1d26cd81ff9f33ee3bd87bb6f6fd2453 100644 (file)
@@ -56,6 +56,14 @@ inductive ex3_3 (A0,A1,A2:Type[0]) (P0,P1,P2:A0→A1→A2→Prop) : Prop ≝
 
 interpretation "multiple existental quantifier (3, 3)" 'Ex P0 P1 P2 = (ex3_3 ? ? ? P0 P1 P2).
 
+(* multiple existental quantifier (4, 2) *)
+
+inductive ex4_2 (A0,A1:Type[0]) (P0,P1,P2,P3:A0→A1→Prop) : Prop ≝
+   | ex4_2_intro: ∀x0,x1. P0 x0 x1 → P1 x0 x1 → P2 x0 x1 → P3 x0 x1 → ex4_2 ? ? ? ? ? ?
+.
+
+interpretation "multiple existental quantifier (4, 2)" 'Ex P0 P1 P2 P3 = (ex4_2 ? ? P0 P1 P2 P3).
+
 (* multiple existental quantifier (4, 3) *)
 
 inductive ex4_3 (A0,A1,A2:Type[0]) (P0,P1,P2,P3:A0→A1→A2→Prop) : Prop ≝
index b8270b6602e48a10c5801be709b2f0c54390fa43..5f747314cc73e2c7b97c13f28089662a06510d72 100644 (file)
@@ -64,6 +64,16 @@ notation < "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break
  non associative with precedence 20
  for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.λ${ident x2}:$T2.$P2) }.
 
+(* multiple existental quantifier (4, 2) *)
+
+notation > "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}.λ${ident x1}.$P0) (λ${ident x0}.λ${ident x1}.$P1) (λ${ident x0}.λ${ident x1}.$P2) (λ${ident x0}.λ${ident x1}.$P3) }.
+
+notation < "hvbox(∃∃ ident x0 , ident x1 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"
+ non associative with precedence 20
+ for @{ 'Ex (λ${ident x0}:$T0.λ${ident x1}:$T1.$P0) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P1) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P2) (λ${ident x0}:$T0.λ${ident x1}:$T1.$P3) }.
+
 (* multiple existental quantifier (4, 3) *)
 
 notation > "hvbox(∃∃ ident x0 , ident x1 , ident x2 break . term 19 P0 break & term 19 P1 break & term 19 P2 break & term 19 P3)"