]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/multi_universal/tuples.ma
some work on extended reduction ...
[helm.git] / matita / matita / lib / turing / multi_universal / tuples.ma
index c565da088aa4f48ab549b3b181ec02b2c6296769..4a7f2b5f2b77ba556cdaddc76fd6b3b91b234769 100644 (file)
@@ -62,10 +62,10 @@ definition is_config : nat → list unialpha → Prop ≝
  only_bits qin ∧ cin ≠ bar ∧ |qin| = S n ∧ t = bar::qin@[cin].
 
 lemma table_to_list: ∀n,l,h,c. is_config n c → 
-  (∃ll,lr.table_TM n l h = ll@c@lr) → 
-    ∃out,t. tuple_encoding n h t = (c@out) ∧ mem ? t l.
+  ∀ll,lr.table_TM n l h = ll@c@lr → 
+    ∃out,lr0,t. lr = out@lr0 ∧ tuple_encoding n h t = (c@out) ∧ mem ? t l.
 #n #l #h #c * #qin * #cin * * * #H1 #H2 #H3 #H4  
- * #ll * #lr lapply ll -ll elim l
+#ll #lr lapply ll -ll elim l
   [>H4 #ll cases ll normalize [|#hd #tl ] #Habs destruct
   |#t1 #othert #Hind #ll >table_TM_cons #Htuple
    cut (S n < |ll@c@lr|)
@@ -85,8 +85,8 @@ lemma table_to_list: ∀n,l,h,c. is_config n c →
      normalize in ⊢ (???%→?); whd in ⊢ (??%?→?); #Htemp 
      lapply (cons_injective_l ????? Htemp) #Hc1
      lapply (cons_injective_r ????? Htemp) -Htemp #Heq2
-     %{(q2@[c2;m])} %{t1} % 
-      [>Ht1 >Heq1 >Hc1 @eq_f >associative_append % 
+     %{(q2@[c2;m])} %{(table_TM n othert h)} %{t1} % 
+      [ %[ // |>Ht1 >Heq1 >Hc1 @eq_f >associative_append % ]
       |%1 %
       ]
     |(* ll not nil *)
@@ -94,8 +94,9 @@ lemma table_to_list: ∀n,l,h,c. is_config n c →
      whd in ⊢ (??%?→?); #Heq destruct (Heq) 
      cases (compare_append … e0) #l *
       [* cases l 
-        [#_ #Htab cases (Hind [ ] (sym_eq … Htab)) #out * #t * #Ht #Hmemt
-         %{out} %{t} % // %2 //
+        [#_ #Htab cases (Hind [ ] (sym_eq … Htab)) 
+         #out * #lr0  * #t * * #Hlr #Ht #Hmemt
+         %{out} %{lr0} %{t} % [% //| %2 //]
         |(* this case is absurd *) 
          #al #tll #Heq1 >H4 #Heq2 @False_ind 
          lapply (cons_injective_l ? bar … Heq2) #Hbar <Hbar in Heq1; #Heq1
@@ -111,8 +112,8 @@ lemma table_to_list: ∀n,l,h,c. is_config n c →
             ]
           ]
         ]
-      |* #Htl #Htab cases (Hind … Htab) #out * #t * #Ht #Hmemt
-       %{out} %{t} % // %2 //
+      |* #Htl #Htab cases (Hind … Htab) #out * #lr0 * #t * * #Hlr #Ht #Hmemt
+       %{out} %{lr0} %{t} % [% // | %2 //]
       ] 
     ]
   ]
@@ -226,59 +227,3 @@ cut (〈q11,a11,m1〉=〈q21,a21,m2〉)
 #Heqout <Heqout in tuplet2; <Heqq <Heqa >tuplet1
 @append_l2_injective %
 qed.
-lemma cfg_in_table_to_tuple: ∀n,l,h,c. is_config n c → 
-  ∀ll,lr.table_TM n l h = ll@c@lr → 
-    ∃out,m,lr0. lr = out@m::lr0 ∧ is_config n (bar::out) ∧ m ≠ bar.
-#n #l #h #c * #qin * #cin * * * #H1 #H2 #H3 #H4  
-#ll #lr lapply ll -ll elim l
-  [>H4 #ll cases ll normalize [|#hd #tl ] #Habs destruct
-  |#t1 #othert #Hind #ll >table_TM_cons #Htuple
-   cut (S n < |ll@c@lr|)
-     [<Htuple >length_append >(length_of_tuple  … (is_tuple … ))
-      /2 by transitive_lt, le_n/] #Hsplit lapply Htuple -Htuple
-   cases (is_tuple … n h t1) #q1 * #c1 * #q2 * #c2 * #m 
-   * * * * * * * #Hq1 #Hq2 #Hc1 #Hc2 #Hm #Hlen1 #Hlen2 
-   whd in ⊢ (???%→?); #Ht1 
-   (* if ll is empty we match the first tuple t1, otherwise
-      we match inside othert *)
-   cases ll
-    [>H4 >Ht1 normalize in ⊢ (???%→?); 
-     >associative_append whd in ⊢ (??%?→?); #Heq destruct (Heq) -Heq
-     >associative_append in e0; #e0
-     lapply (append_l1_injective  … e0) [>H3 @Hlen1] #Heq1
-     lapply (append_l2_injective  … e0) [>H3 @Hlen1]
-     normalize in ⊢ (???%→?); whd in ⊢ (??%?→?); #Htemp 
-     lapply (cons_injective_l ????? Htemp) #Hc1
-     lapply (cons_injective_r ????? Htemp) -Htemp #Heq2
-     %{(q2@[c2])} %{m} %{(table_TM n othert h)} % // %
-     [ <Heq2 >associative_append >associative_append % | %{q2} %{c2} % // % // % // ] 
-    |(* ll not nil *)
-     #b #tl >Ht1 normalize in ⊢ (???%→?); 
-     whd in ⊢ (??%?→?); #Heq destruct (Heq) 
-     cases (compare_append … e0) #l *
-      [* cases l 
-        [#_ #Htab cases (Hind [ ] (sym_eq … Htab)) #out * #m0 * #lr0 * * #Hlr #Hcfg #Hm0
-         %{out} %{m0} %{lr0} % // % //
-        |(* this case is absurd *) 
-         #al #tll #Heq1 >H4 #Heq2 @False_ind 
-         lapply (cons_injective_l ? bar … Heq2) #Hbar <Hbar in Heq1; #Heq1
-         @(absurd (mem ? bar (q1@(c1::q2@[c2; m]))))
-          [>Heq1 @mem_append_l2 %1 //
-          |% #Hmembar cases (mem_append ???? Hmembar) -Hmembar
-            [#Hmembar lapply(Hq1 bar Hmembar) normalize #Habs destruct (Habs)
-            |* [#Habs @absurd //]
-             #Hmembar cases (mem_append ???? Hmembar) -Hmembar
-              [#Hmembar lapply(Hq2 bar Hmembar) normalize #Habs destruct (Habs)
-              |* [#Habs @absurd //] #Hmembar @(absurd ?? Hm) @sym_eq @mem_single //
-              ]
-            ]
-          ]
-        ]
-      |* #Htl #Htab cases (Hind … Htab) #out * #m0 * #lr0 * * #Hlr #Hcfg #Hm0
-        %{out} %{m0} %{lr0} % // % //
-      ] 
-    ]
-  ]
-qed.
-