]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/universal/tuples.ma
Msg reporting via HLogger in the error window made asynchronous => speedup.
[helm.git] / matita / matita / lib / turing / universal / tuples.ma
index 18da8b92eafa0531f4d657aa50b755c4d689afd0..0faa4eece58844e8925519b42899d97c2b799be3 100644 (file)
@@ -19,7 +19,8 @@ include "turing/universal/marks.ma".
 definition STape ≝ FinProd … FSUnialpha FinBool.
 
 definition only_bits ≝ λl.
-  ∀c.memb STape c l = true → is_bit (\fst c) = true.
+  ∀c.memb STape c l
+   = true → is_bit (\fst c) = true.
 
 definition only_bits_or_nulls ≝ λl.
   ∀c.memb STape c l = true → bit_or_null (\fst c) = true.
@@ -50,8 +51,8 @@ lemma bit_or_null_not_bar: ∀d. bit_or_null d = true → is_bar d = false.
 qed.
 
 definition mk_tuple ≝ λqin,cin,qout,cout,mv.
-  qin @ cin :: 〈comma,false〉:: qout @ cout :: 〈comma,false〉 :: [mv].
-  
+  〈bar,false〉 :: qin @ cin :: 〈comma,false〉:: qout @ cout :: 〈comma,false〉 :: [mv].
+
 (* by definition, a tuple is not marked *)
 definition tuple_TM : nat → list STape → Prop ≝ 
  λn,t.∃qin,cin,qout,cout,mv.
@@ -64,43 +65,107 @@ definition tuple_TM : nat → list STape → Prop ≝
  
 inductive table_TM (n:nat) : list STape → Prop ≝ 
 | ttm_nil  : table_TM n [] 
-| ttm_cons : ∀t1,T.tuple_TM n t1 → table_TM n T → table_TM n (t1@〈bar,false〉::T).
+| ttm_cons : ∀t1,T.tuple_TM n t1 → table_TM n T → table_TM n (t1@T).
 
-inductive match_in_table (qin:list STape) (cin: STape) 
+inductive match_in_table (n:nat) (qin:list STape) (cin: STape) 
                          (qout:list STape) (cout:STape) (mv:STape) 
 : list STape → Prop ≝ 
 | mit_hd : 
    ∀tb.
-   match_in_table qin cin qout cout mv 
-     (mk_tuple qin cin qout cout mv @〈bar,false〉::tb)
+   tuple_TM n (mk_tuple qin cin qout cout mv) → 
+   match_in_table n qin cin qout cout mv 
+     (mk_tuple qin cin qout cout mv @tb)
 | mit_tl :
    ∀qin0,cin0,qout0,cout0,mv0,tb.
-   match_in_table qin cin qout cout mv tb → 
-   match_in_table qin cin qout cout mv  
-     (mk_tuple qin0 cin0 qout0 cout0 mv0@〈bar,false〉::tb).
+   tuple_TM n (mk_tuple qin0 cin0 qout0 cout0 mv0) → 
+   match_in_table n qin cin qout cout mv tb → 
+   match_in_table n qin cin qout cout mv  
+     (mk_tuple qin0 cin0 qout0 cout0 mv0@tb).
      
-axiom append_l1_injective : 
-  ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l1 = l2.
-axiom append_l2_injective : 
-  ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l3 = l4.
-axiom cons_injective_l : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → a1 = a2.
-axiom cons_injective_r : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → l1 = l2.
-axiom tuple_len : ∀n,t.tuple_TM n t → |t| = 2*n+5.
+axiom tuple_len : ∀n,t.tuple_TM n t → |t| = 2*n+6.
 axiom append_eq_tech1 :
-  ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → |l1| < |l3| → ∃la:list A.l1@a::la = l3.
+  ∀A,l1,l2,l3,l4.l1@l2 = l3@l4 → |l1| < |l3| → ∃la:list A.l1@la = l3.
 axiom append_eq_tech2 :
   ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false → ∃la:list A.l3 = l1@a::la.
 (*axiom list_decompose_cases : 
   ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → ∃la,lb:list A.l3 = la@a::lb ∨ l4 = la@a::lb.
 axiom list_decompose_l :
   ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l4 = false → 
-  ∃la,lb.l2 = la@lb ∧ l3 = l1@a::la.
+  ∃la,lb.l2 = la@lb ∧ l3 = l1@a::la.*)
 axiom list_decompose_r :
   ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → memb A a l3 = false → 
   ∃la,lb.l1 = la@lb ∧ l4 = lb@a::l2.
-axiom list_decompose_memb :
+(*axiom list_decompose_memb :
   ∀A,l1,l2,l3,l4,a.l1@a::l2 = l3@l4 → |l1| < |l3| → memb A a l3 = true.*)
 
+lemma table_invert_r : ∀n,t,T.
+  tuple_TM n t → table_TM n (t@T) → table_TM n T.
+#n #t #T #Htuple #Htable inversion Htable
+[ cases Htuple #qin * #cin * #qout * #cout * #mv * #_ #Ht >Ht
+  normalize #Hfalse destruct (Hfalse)
+| #t0 #T0 #Htuple0 #Htable0 #_ #Heq 
+  lapply (append_l2_injective ?????? Heq)
+  [ >(tuple_len … Htuple) >(tuple_len … Htuple0) % ]
+  -Heq #Heq destruct (Heq) // ]
+qed.
+
+lemma match_in_table_to_tuple :
+  ∀n,T,qin,cin,qout,cout,mv.
+  match_in_table n qin cin qout cout mv T → table_TM n T → 
+  tuple_TM n (mk_tuple qin cin qout cout mv).
+#n #T #qin #cin #qout #cout #mv #Hmatch elim Hmatch
+[ //
+| #qin0 #cin0 #qout0 #cout0 #mv0 #tb #Htuple #Hmatch #IH #Htable
+  @IH @(table_invert_r ???? Htable) @Htuple
+]
+qed.
+
+lemma match_in_table_append :
+  ∀n,T,qin,cin,qout,cout,mv,t.
+  tuple_TM n t → 
+  match_in_table n qin cin qout cout mv (t@T) → 
+  t = mk_tuple qin cin qout cout mv ∨ match_in_table n qin cin qout cout mv T.
+#n #T #qin #cin #qout #cout #mv #t #Ht #Hmatch inversion Hmatch
+[ #T0 #H #H1 % >(append_l1_injective … H1) //
+  >(tuple_len … Ht) >(tuple_len … H) %
+| #qin0 #cin0 #qout0 #cout0 #mv0 #T0 #H #H1 #_ #H2 %2
+  >(append_l2_injective … H2) // >(tuple_len … Ht) >(tuple_len … H) %
+]
+qed.
+
+lemma generic_match_to_match_in_table_tech : 
+  ∀n,t,T0,T1,T2.tuple_TM n t → table_TM n (T1@〈bar,false〉::T2) → 
+   t@T0 = T1@〈bar,false〉::T2 → T1 = [] ∨ ∃T3.T1 = t@T3.
+#n #t #T0 #T1 #T2 #Ht cases T1
+[ #_ #_ % %
+| normalize #c #T1c #Htable #Heq %2
+  cases Ht in Heq; #qin * #cin * #qout * #cout * #mv **********
+  #Hqin1 #Hqout1 #Hqin2 #Hqout2 #Hcin #Hcout #Hmv #Hcoutmv #Hqinlen #Hqoutlen
+  #Heqt >Heqt whd in ⊢ (??%%→?); #Ht lapply (cons_injective_r ????? Ht)
+  #Ht' cases (list_decompose_r STape … (sym_eq … Ht') ?)
+  [ #la * #lb * #HT1c #HT0 %{lb} >HT1c @(eq_f2 ??? (append ?) (c::la)) //
+    >HT0 in Ht'; >HT1c >associative_append in ⊢ (???%→?); #Ht'
+    <(append_l1_injective_r … Ht') // <(cons_injective_l ????? Ht) %
+  |@(noteq_to_eqnot ? true) @(not_to_not … not_eq_true_false) #Hbar @sym_eq 
+   cases (memb_append … Hbar) -Hbar #Hbar
+    [@(Hqin2 … Hbar) 
+    |cases (orb_true_l … Hbar) -Hbar
+      [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcin
+      |whd in ⊢ ((??%?)→?); #Hbar cases (memb_append … Hbar) -Hbar #Hbar
+        [@(Hqout2 … Hbar)
+        |cases (orb_true_l … Hbar) -Hbar
+          [#Hbar lapply (\P Hbar) -Hbar #Hbar destruct (Hbar) @Hcout
+          |#Hbar cases (orb_true_l … Hbar) -Hbar 
+            [whd in ⊢ ((??%?)→?); #Hbar @Hbar
+            |#Hbar lapply (memb_single … Hbar) -Hbar #Hbar destruct (Hbar) @Hmv
+            ]
+          ]
+        ]
+      ]
+    ]
+  ]
+qed.
+    
 lemma generic_match_to_match_in_table :
   ∀n,T.table_TM n T → 
   ∀qin,cin,qout,cout,mv.|qin| = n → |qout| = n → 
@@ -108,23 +173,29 @@ lemma generic_match_to_match_in_table :
   bit_or_null (\fst cin) = true → bit_or_null (\fst cout) = true → 
   bit_or_null (\fst mv) = true →  
   ∀t1,t2.
-  T = (t1@qin@cin::〈comma,false〉::qout@cout::〈comma,false〉::[mv])@t2 → 
-  match_in_table qin cin qout cout mv T.
+  T = (t1@〈bar,false〉::qin@cin::〈comma,false〉::qout@cout::〈comma,false〉::[mv])@t2 → 
+  match_in_table qin cin qout cout mv T.
 #n #T #Htable #qin #cin #qout #cout #mv #Hlenqin #Hlenqout
 #Hqinbits #Hqoutbits #Hcin #Hcout #Hmv
 elim Htable
-[ #t1 #t2 <associative_append cases (t1@qin) normalize
-  [ #Hfalse destruct (Hfalse) | #c0 #t0 #Hfalse destruct (Hfalse) ]
-| #tuple #T0 * #qin0 * #cin0 * #qout0 * #cout0 * #mv0
+[ * [ #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
+    | #c0 #t1 #t2 normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ]
+| #tuple #T0 #H1 #Htable0#IH #t1 #t2 #HT cases H1 #qin0 * #cin0 * #qout0 * #cout0 * #mv0
   * * * * * * * * * *
   #Hqin0marks #Hqout0marks #Hqin0bits #Hqout0bits #Hcin0 #Hcout0 #Hmv0 #Hcout0mv0
-  #Hlenqin0 #Hlenqout0 #Htuple #Htable0 #IH #t1 #t2 #HT
-  cases t1 in HT;
-  [ >Htuple normalize in ⊢ (??%%→?);
+  #Hlenqin0 #Hlenqout0 #Htuple 
+  lapply (generic_match_to_match_in_table_tech n ? T0 t1 
+           (qin@cin::〈comma,false〉::qout@[cout;〈comma,false〉;mv]@t2) H1) #Htmp
+  >Htuple in H1; #H1 
+  lapply (ttm_cons … T0 H1 Htable0) <Htuple in ⊢ (%→?); >HT
+  >associative_append normalize >associative_append normalize
+  >associative_append #Htable cases (Htmp Htable ?)
+  [ #Ht1 >Htuple in HT; >Ht1 normalize in ⊢ (??%%→?);
     >associative_append >associative_append #HT
     cut (qin0 = qin ∧ (〈cin0,false〉 = cin ∧ (qout0 = qout ∧ 
-         (〈cout0,false〉 = cout ∧ (〈mv0,false〉 = mv ∧ 〈bar,false〉::T0 = t2)))))
-    [ lapply (append_l1_injective … HT) [ >Hlenqin @Hlenqin0 ]
+         (〈cout0,false〉 = cout ∧ (〈mv0,false〉 = mv ∧ T0 = t2)))))
+    [ lapply (cons_injective_r ????? HT) -HT #HT 
+      lapply (append_l1_injective … HT) [ >Hlenqin @Hlenqin0 ]
       #Hqin % [ @Hqin ] -Hqin
       lapply (append_l2_injective … HT) [ >Hlenqin @Hlenqin0 ] -HT #HT
       lapply (cons_injective_l ????? HT) #Hcin % [ @Hcin ] -Hcin
@@ -140,33 +211,38 @@ elim Htable
       lapply (cons_injective_l ????? HT) #Hmv % [ @Hmv ] -Hmv
       @(cons_injective_r ????? HT) ]
     -HT * #Hqin * #Hcin * #Hqout * #Hcout * #Hmv #HT0
-    >(?:qin0@(〈cin0,false〉::〈comma,false〉::qout0@[〈cout0,false〉;〈comma,false〉;〈mv0,false〉])@〈bar,false〉::T0
-        = mk_tuple qin cin qout cout mv@〈bar,false〉::T0)
-    [|>Hqin >Hqout >Hcin >Hcout >Hmv normalize >associative_append >associative_append
-       normalize >associative_append % ]
-    %
-  | #c0 #cs0 #HT cut (∃cs1.c0::cs0 = tuple@〈bar,false〉::cs1)
-    [ cases (append_eq_tech1 ?????? HT ?)
-      [ -HT #ta #Hta cases (append_eq_tech2 … Hta ?)
-        [ -Hta #tb #Htb %{tb} @Htb 
-        | @daemon ]
-      | @le_S_S >length_append >(plus_n_O (|tuple|)) >commutative_plus @le_plus
-        [ @le_O_n
-        | >Htuple normalize >length_append >length_append @le_plus [ >Hlenqin >Hlenqin0 % ]
-          @le_S_S @le_S_S >length_append >length_append @le_plus [ >Hlenqout >Hlenqout0 % ] %] ] 
-    ]
-    * #cs1 #Hcs1 >Hcs1 in HT; >associative_append >associative_append #HT
-    lapply (append_l2_injective … HT) // -HT #HT
-    lapply (cons_injective_r ????? HT) -HT
-    <associative_append #HT >Htuple %2 @(IH ?? HT)
-  ]
+    >(?:〈bar,false〉::qin0@(〈cin0,false〉::〈comma,false〉::qout0@
+        [〈cout0,false〉;〈comma,false〉;〈mv0,false〉])@T0 = tuple@T0)
+    [ >Htuple >Hqin >Hqout >Hcin >Hcout >Hmv % //
+    | >Htuple normalize >associative_append normalize >associative_append
+      normalize >associative_append % ]
+  | * #T3 #HT3 >HT3 in HT; >associative_append; >associative_append #HT
+    lapply (append_l2_injective … HT) // -HT #HT %2 //
+    @(IH T3 t2) >HT >associative_append %
+  |>HT >associative_append normalize >associative_append normalize
+   >associative_append % ]
 ]
 qed.
 
+(*
+lemma table_invert_l : ∀n,T0,qin,cin,qout,cout,mv.
+  table_TM n (mk_tuple qin cin qout cout mv@〈bar,false〉::T0) → 
+  tuple_TM n (mk_tuple qin cin qout cout mv).
+#n #T #qin #cin #qout #cout #mv #HT inversion HT
+[ change with (append ???) in ⊢ (??(??%?)?→?);cases qin [ #Hfalse | #t0 #ts0 #Hfalse] normalize in Hfalse; destruct (Hfalse)
+| #t0 #T0 #Ht0 #HT0 #_
+
+  
+lemma table_invert_r : ∀n,T0,qin,cin,qout,cout,mv.
+  table n (mk_tuple qin cin qout cout mv@〈bar,false〉::T0) → table n T0. 
+*)
+
 lemma no_grids_in_tuple : ∀n,l.tuple_TM n l → no_grids l.
 #n #l * #qin * #cin * #qout * #cout * #mv * * * * * * * * * *
 #_ #_ #Hqin #Hqout #Hcin #Hcout #Hmv #_ #_ #_ #Hl >Hl
-#c #Hc normalize in Hc; cases (memb_append … Hc) -Hc #Hc
+#c #Hc cases (orb_true_l … Hc) -Hc #Hc
+[ >(\P Hc) %
+| cases (memb_append … Hc) -Hc #Hc
 [ @bit_not_grid @(Hqin … Hc)
 | cases (orb_true_l … Hc) -Hc #Hc 
 [ change with (c == 〈cin,false〉 = true) in Hc; >(\P Hc) @bit_or_null_not_grid //
@@ -185,7 +261,9 @@ qed.
 lemma no_marks_in_tuple : ∀n,l.tuple_TM n l → no_marks l.
 #n #l * #qin * #cin * #qout * #cout * #mv * * * * * * * * * *
 #Hqin #Hqout #_ #_ #_ #_ #_ #_ #_ #_ #Hl >Hl
-#c #Hc normalize in Hc; cases (memb_append … Hc) -Hc #Hc
+#c #Hc cases (orb_true_l … Hc) -Hc #Hc
+[ >(\P Hc) %
+| cases (memb_append … Hc) -Hc #Hc
 [ @(Hqin … Hc)
 | cases (orb_true_l … Hc) -Hc #Hc 
 [ change with (c == 〈cin,false〉 = true) in Hc; >(\P Hc) %
@@ -207,9 +285,7 @@ lemma no_grids_in_table: ∀n.∀l.table_TM n l → no_grids l.
   |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_grids_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx
    cases (memb_append … Hx) -Hx #Hx
    [ @(Ht1 … Hx)
-   | cases (orb_true_l … Hx) -Hx #Hx
-     [ >(\P Hx) %
-     | @(IH … Hx) ] ] ]
+   | @(IH … Hx) ] ]
 qed.
 
 lemma no_marks_in_table: ∀n.∀l.table_TM n l → no_marks l.
@@ -218,9 +294,7 @@ lemma no_marks_in_table: ∀n.∀l.table_TM n l → no_marks l.
   |#t1 #t2 #Ht1 #Ht2 #IH lapply (no_marks_in_tuple … Ht1) -Ht1 #Ht1 #x #Hx
    cases (memb_append … Hx) -Hx #Hx
    [ @(Ht1 … Hx)
-   | cases (orb_true_l … Hx) -Hx #Hx
-     [ >(\P Hx) %
-     | @(IH … Hx) ] ] ]
+   | @(IH … Hx) ] ] 
 qed.      
           
 axiom last_of_table: ∀n,l,b.¬ table_TM n (l@[〈bar,b〉]).
@@ -256,7 +330,7 @@ definition bar_or_grid ≝ λc:STape.is_bar (\fst c) ∨ is_grid (\fst c).
 definition mark_next_tuple ≝ 
   seq ? (adv_to_mark_r ? bar_or_grid)
      (ifTM ? (test_char ? (λc:STape.is_bar (\fst c)))
-       (move_right_and_mark ?) (nop ?) 1).
+       (move_right_and_mark ?) (nop ?) tc_true).
 
 definition R_mark_next_tuple ≝ 
   λt1,t2.
@@ -271,6 +345,8 @@ definition R_mark_next_tuple ≝
     ∨
     (no_bars rs1 ∧ t2 = midtape ? (reverse ? rs1@c::ls) 〈grid,false〉 rs2).
      
+axiom daemon :∀P:Prop.P.
+
 axiom tech_split :
   ∀A:DeqSet.∀f,l.
    (∀x.memb A x l = true → f x = false) ∨
@@ -282,7 +358,7 @@ theorem sem_mark_next_tuple :
   Realize ? mark_next_tuple R_mark_next_tuple.
 #intape 
 lapply (sem_seq ? (adv_to_mark_r ? bar_or_grid)
-         (ifTM ? (test_char ? (λc:STape.is_bar (\fst c))) (move_right_and_mark ?) (nop ?) 1) ????)
+         (ifTM ? (test_char ? (λc:STape.is_bar (\fst c))) (move_right_and_mark ?) (nop ?) tc_true) ????)
 [@sem_if [5: // |6: @sem_move_right_and_mark |7: // |*:skip]
 | //
 |||#Hif cases (Hif intape) -Hif
@@ -299,7 +375,7 @@ lapply (sem_seq ? (adv_to_mark_r ? bar_or_grid)
        | -Hta #Hta cases Hright
          [ * #tb * whd in ⊢ (%→?); #Hcurrent
            @False_ind cases (Hcurrent 〈grid,false〉 ?)
-           [ normalize #Hfalse destruct (Hfalse)
+           [ normalize in ⊢ (%→?); #Hfalse destruct (Hfalse)
            | >Hta % ]
          | * #tb * whd in ⊢ (%→?); #Hcurrent
            cases (Hcurrent 〈grid,false〉 ?)
@@ -449,6 +525,7 @@ generalize in match Hc; generalize in match Hl2; cases l2
    ]
 qed.
 *)
+
 definition init_current ≝ 
   seq ? (adv_to_mark_l ? (is_marked ?))
     (seq ? (clear_mark ?)
@@ -503,29 +580,32 @@ definition match_tuple_step ≝
     (nop ?) tc_true.
 
 definition R_match_tuple_step_true ≝ λt1,t2.
-  ∀ls,c,l1,l2,c1,l3,l4,rs,n.
-  bit_or_null c = true → only_bits_or_nulls l1 → no_marks l1 (* → no_grids l2 *) → bit_or_null c1 = true →
-  only_bits_or_nulls l3 → n = |l1| → |l1| = |l3| →
-  table_TM (S n) (l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::l4) → 
-  t1 = midtape STape (〈grid,false〉::ls) 〈c,true〉 
-         (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs) → 
-  (* facciamo match *)
-  (〈c,false〉::l1 = 〈c1,false〉::l3 ∧
-  t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉
-        (l2@〈bar,false〉::〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs))
-  ∨
-  (* non facciamo match e marchiamo la prossima tupla *)
-  ((〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧
-   ∃c2,l5,l6.l4 = l5@〈bar,false〉::〈c2,false〉::l6 ∧
-   (* condizioni su l5 l6 l7 *)
-   t2 = midtape STape (〈grid,false〉::ls) 〈c,true〉 
-         (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::
-          l5@〈bar,false〉::〈c2,true〉::l6@〈grid,false〉::rs))
-  ∨  
-  (* non facciamo match e non c'è una prossima tupla:
-     non specifichiamo condizioni sul nastro di output, perché
-     non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
-  (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉)).  
+  ∀ls,cur,rs.t1 = midtape STape ls cur rs → 
+  \fst cur ≠ grid ∧ 
+  (∀ls0,c,l1,l2,c1,l3,l4,rs0,n.
+   only_bits_or_nulls l1 → no_marks l1 (* → no_grids l2 *) → 
+   bit_or_null c = true → bit_or_null c1 = true →
+   only_bits_or_nulls l3 → S n = |l1| → |l1| = |l3| →
+   table_TM (S n) (l2@〈c1,false〉::l3@〈comma,false〉::l4) → 
+   ls = 〈grid,false〉::ls0 → cur = 〈c,true〉 → 
+   rs = l1@〈grid,false〉::l2@〈c1,true〉::l3@〈comma,false〉::l4@〈grid,false〉::rs0 → 
+   (* facciamo match *)
+   (〈c,false〉::l1 = 〈c1,false〉::l3 ∧
+   t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
+         (l2@〈c1,false〉::l3@〈comma,true〉::l4@〈grid,false〉::rs0))
+   ∨
+   (* non facciamo match e marchiamo la prossima tupla *)
+   (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧
+    ∃c2,l5,l6.l4 = l5@〈bar,false〉::〈c2,false〉::l6 ∧
+    (* condizioni su l5 l6 l7 *)
+    t2 = midtape STape (〈grid,false〉::ls0) 〈c,true〉 
+          (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::
+           l5@〈bar,false〉::〈c2,true〉::l6@〈grid,false〉::rs0))
+   ∨  
+   (* non facciamo match e non c'è una prossima tupla:
+      non specifichiamo condizioni sul nastro di output, perché
+      non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
+   (〈c,false〉::l1 ≠ 〈c1,false〉::l3 ∧ no_bars l4 ∧ current ? t2 = Some ? 〈grid,true〉)).  
   
 definition R_match_tuple_step_false ≝ λt1,t2.
   ∀ls,c,rs.t1 = midtape STape ls c rs → is_grid (\fst c) = true ∧ t2 = t1.
@@ -551,7 +631,7 @@ axiom tech_split2 : ∀A,l1,l2,l3,l4,x.
 axiom injective_append : ∀A,l.injective … (λx.append A x l).
 
 lemma sem_match_tuple_step: 
-    accRealize ? match_tuple_step (inr … (inl … (inr … 0))) 
+    accRealize ? match_tuple_step (inr … (inl … (inr … start_nop))) 
     R_match_tuple_step_true R_match_tuple_step_false.
 @(acc_sem_if_app … (sem_test_char ? (λc:STape.¬ is_grid (\fst c))) …
   (sem_seq … sem_compare
@@ -565,20 +645,22 @@ lemma sem_match_tuple_step:
  2:#t1 #t2 #t3 whd in ⊢ (%→?); #H #H1 whd #ls #c #rs #Ht1
   cases (H c ?) [2: >Ht1 %] #Hgrid #Heq %
     [@injective_notb @Hgrid | <Heq @H1]
-|#tapea #tapeout #tapeb whd in ⊢ (%→?); #Htapea
- * #tapec * #Hcompare #Hor 
- #ls #c #l1 #l2 #c1 #l3 #l4 #rs #n #Hc #Hl1bars #Hl1marks #Hc1 #Hl3 #eqn
- #eqlen #Htable #Htapea1 cases (Htapea 〈c,true〉 ?) >Htapea1 [2:%]
- #notgridc -Htapea -Htapea1 -tapea #Htapeb  
- cases (Hcompare … Htapeb) -Hcompare -Htapeb * #_ #_ #Hcompare 
- cases (Hcompare c c1 l1 l3 (l2@[〈bar,false〉]) (l4@〈grid,false〉::rs) eqlen Hl1bars Hl3 Hl1marks … (refl …) Hc ?)  
+|#tapea #tapeout #tapeb whd in ⊢ (%→?); #Hcur
+ * #tapec * whd in ⊢ (%→?); #Hcompare #Hor 
+ #ls #cur #rs #Htapea >Htapea in Hcur; #Hcur cases (Hcur ? (refl ??)) 
+ -Hcur #Hcur #Htapeb %
+ [ % #Hfalse >Hfalse in Hcur; normalize #Hfalse1 destruct (Hfalse1)]
+ #ls0 #c #l1 #l2 #c1 #l3 #l4 #rs0 #n #Hl1bitnull #Hl1marks #Hc #Hc1 #Hl3 #eqn
+ #eqlen #Htable #Hls #Hcur #Hrs -Htapea >Hls in Htapeb; >Hcur >Hrs #Htapeb
+ cases (Hcompare … Htapeb) -Hcompare -Htapeb * #_ #_ #Hcompare
+ cases (Hcompare c c1 l1 l3 l2 (l4@〈grid,false〉::rs0) eqlen Hl1bitnull Hl3 Hl1marks … (refl …) Hc ?)  
  -Hcompare 
-   [* #Htemp destruct (Htemp) #Htapec %1 % [%]
+   [* #Htemp destruct (Htemp) #Htapec %1 % [%]
     >Htapec in Hor; -Htapec *
      [2: * #t3 * whd in ⊢ (%→?); #H @False_ind
       cases (H … (refl …)) whd in ⊢ ((??%?)→?); #H destruct (H)
      |* #taped * whd in ⊢ (%→?); #Htaped cases (Htaped ? (refl …)) -Htaped *
-      #Htaped whd in ⊢ (%→?); #Htapeout >Htapeout >Htaped >associative_append
+      #Htaped whd in ⊢ (%→?); #Htapeout >Htapeout >Htaped
       %
      ]
    |* #la * #c' * #d' * #lb * #lc * * * #H1 #H2 #H3 #Htapec 
@@ -587,7 +669,7 @@ lemma sem_match_tuple_step:
         [@(not_to_not …H1) normalize #H destruct % 
         |#x #tl @not_to_not normalize #H destruct // 
         ]
-      ] #Hnoteq %2
+      ] #Hnoteq
     cut (bit_or_null d' = true) 
       [cases la in H3;
         [normalize in ⊢ (%→?); #H destruct //
@@ -608,13 +690,13 @@ lemma sem_match_tuple_step:
           * #d * #b * * * #Heq1 @False_ind 
           cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut 
           >Hcut in Htable; >H3 >associative_append
-          normalize >Heq1 >Hcut <associative_append >Hcut 
+          normalize >Heq1 <associative_append >Hcut
           <associative_append #Htable @(absurd … Htable) 
-          @last_of_table    
+          @last_of_table
          |(* rs4 not empty *)
           * #d2 #b2 #rs3' * #d  * #b * * * #Heq1 #Hnobars
-          cut (memb STape 〈d2,b2〉 (l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::l4) = true)
-            [@memb_append_l2 @memb_cons
+          cut (memb STape 〈d2,b2〉 (l2@〈c1,false〉::l3@〈comma,false〉::l4) = true)
+            [@memb_append_l2
              cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut
              >Hcut >H3 >associative_append @memb_append_l2 
              @memb_cons >Heq1 @memb_append_l2 @memb_cons @memb_hd] #d2intable
@@ -636,126 +718,111 @@ lemma sem_match_tuple_step:
             >Htapeg -Htapeg
             (* init_current *)
              whd in ⊢ (%→?); #Htapeout
-            %1 cases (some_option_hd ? (reverse ? (reverse ? la)) 〈c',true〉)
-            * #c00 #b00 #Hoption
-            lapply 
-             (Htapeout (reverse ? rs3 @〈d',false〉::reverse ? la@reverse ? (l2@[〈bar,false〉])@(〈grid,false〉::reverse ? lb))
-             c' (reverse ? la) false ls bar (〈d2,true〉::rs3'@〈grid,false〉::rs) c00 b00 ?????) -Htapeout
-              [whd in ⊢ (??(??%??)?); @eq_f3 [2:%|3: %]
-               >associative_append 
-               generalize in match (〈c',true〉::reverse ? la@〈grid,false〉::ls); #l
-               whd in ⊢ (???(???%)); >associative_append >associative_append % 
-              |>reverse_cons @Hoption
-              |cases la in H2; 
-                [normalize in ⊢ (%→?); #Htemp destruct (Htemp) 
-                 @injective_notb @notgridc
-                |#x #tl normalize in ⊢ (%→?); #Htemp destruct (Htemp)
-                 @bit_or_null_not_grid @(Hl1bars 〈c',false〉) @memb_append_l2 @memb_hd
-                ]
-              |cut (only_bits_or_nulls (la@(〈c',false〉::lb)))
-                [<H2 whd #c0 #Hmemb cases (orb_true_l … Hmemb)
-                  [#eqc0 >(\P eqc0) @Hc |@Hl1bars]
-                |#Hl1' #x #Hx @bit_or_null_not_grid @Hl1'
-                 @memb_append_l1 @daemon
-                ]
-              |@daemon
-              |>reverse_append >reverse_cons >reverse_reverse
-               >reverse_append >reverse_reverse
-               >reverse_cons >reverse_append >reverse_reverse
-               >reverse_append >reverse_cons >reverse_reverse
-               >reverse_reverse
-               #Htapeout % [@Hnoteq]
-               @(ex_intro … d2)
-               cut (∃rs32.rs3 = lc@〈comma,false〉::rs32) 
-               [ (*cases (tech_split STape (λc.c == 〈bar,false〉) l4)
-                 [
-                 | * #l41 * * #cbar #bfalse * #l42 * * #Hbar #Hl4 #Hl41
-                   @(ex_intro ?? l41) >Hl4 in Heq1; #Heq1
-               
-               cut (sublist … lc l3)
-                 [ #x #Hx cases la in H3;
-                   [ normalize #H3 destruct (H3) @Hx
-                   | #p #la' normalize #Hla' destruct (Hla')
-                     @memb_append_l2 @memb_cons @Hx ] ] #Hsublist*)
-               @daemon]
-               * #rs32 #Hrs3
-               (* cut 
-               (〈c1,false〉::l3@〈comma,false〉::l4= la@〈d',false〉::rs3@〈bar,false〉::〈d2,b2〉::rs3')
-               [@daemon] #Hcut *)
-               cut (l4=rs32@〈bar,false〉::〈d2,false〉::rs3')
-               [ >Hrs3 in Heq1; @daemon ] #Hl4
-               @(ex_intro … rs32) @(ex_intro … rs3') %
-                 [@Hl4
-                 |>Htapeout @eq_f2
-                   [@daemon
+             cases (some_option_hd ? (reverse ? (reverse ? la)) 〈c',true〉)
+             * #c00 #b00 #Hoption
+             lapply 
+              (Htapeout (reverse ? rs3 @〈d',false〉::reverse ? la@reverse ? l2@(〈grid,false〉::reverse ? lb))
+              c' (reverse ? la) false ls0 bar (〈d2,true〉::rs3'@〈grid,false〉::rs0) c00 b00 ?????) -Htapeout
+               [whd in ⊢ (??(??%??)?); @eq_f3 [2:%|3: %]
+                >associative_append 
+                 generalize in match (〈c',true〉::reverse ? la@〈grid,false〉::ls0); #l
+                whd in ⊢ (???(???%)); >associative_append >associative_append % 
+               |>reverse_cons @Hoption
+               |cases la in H2; 
+                 [normalize in ⊢ (%→?); #Htemp destruct (Htemp) 
+                  @bit_or_null_not_grid @Hc
+                 |#x #tl normalize in ⊢ (%→?); #Htemp destruct (Htemp)
+                  @bit_or_null_not_grid @(Hl1bitnull 〈c',false〉) @memb_append_l2 @memb_hd
+                 ]
+               |cut (only_bits_or_nulls (la@(〈c',false〉::lb)))
+                 [<H2 whd #c0 #Hmemb cases (orb_true_l … Hmemb)
+                   [#eqc0 >(\P eqc0) @Hc |@Hl1bitnull]
+                 |#Hl1' #x #Hx @bit_or_null_not_grid @Hl1'
+                  @memb_append_l1 @daemon
+                 ]
+               |@daemon] #Htapeout % %2 % //
+            @(ex_intro … d2)
+            cut (∃rs32.rs3 = lc@〈comma,false〉::rs32) 
+                 [ (*cases (tech_split STape (λc.c == 〈bar,false〉) l4)
+                  [
+                  | * #l41 * * #cbar #bfalse * #l42 * * #Hbar #Hl4 #Hl41
+                    @(ex_intro ?? l41) >Hl4 in Heq1; #Heq1
+                
+                cut (sublist … lc l3)
+                  [ #x #Hx cases la in H3;
+                    [ normalize #H3 destruct (H3) @Hx
+                    | #p #la' normalize #Hla' destruct (Hla')
+                      @memb_append_l2 @memb_cons @Hx ] ] #Hsublist*)
+                @daemon]
+                * #rs32 #Hrs3
+                (* cut 
+                (〈c1,false〉::l3@〈comma,false〉::l4= la@〈d',false〉::rs3@〈bar,false〉::〈d2,b2〉::rs3')
+                [@daemon] #Hcut *)
+                cut (l4=rs32@〈bar,false〉::〈d2,false〉::rs3')
+                [ >Hrs3 in Heq1; @daemon ] #Hl4
+                @(ex_intro … rs32) @(ex_intro … rs3') % [@Hl4]
+                >Htapeout @eq_f2
+                   [(* by Hoption, H2 *) @daemon
                    |(*>Hrs3 *)>append_cons
-                    > (?:l1@〈grid,false〉::l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs
-                        = (l1@〈grid,false〉::l2@〈bar,false〉::〈c1,false〉::l3@〈comma,false〉::rs32@[〈bar,false〉])@〈d2,true〉::rs3'@〈grid,false〉::rs)
+                    > (?:l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs
+                        = (l1@〈grid,false〉::l2@〈c1,false〉::l3@〈comma,false〉::rs32@[〈bar,false〉])@〈d2,true〉::rs3'@〈grid,false〉::rs)
                     [|>associative_append normalize 
                       >associative_append normalize
                       >associative_append normalize
                       >associative_append normalize
                        % ]
-                    @eq_f2 [|%]
-                    @(injective_append … (〈d2,false〉::rs3'))
-                    >(?:(la@[〈c',false〉])@((((lb@[〈grid,false〉])@l2@[〈bar,false〉])@la)@[〈d',false〉])@rs3
-                       =((la@〈c',false〉::lb)@([〈grid,false〉]@l2@[〈bar,false〉]@la@[〈d',false〉]@rs3)))
+                    >reverse_append >reverse_append >reverse_cons
+                    >reverse_reverse >reverse_cons >reverse_reverse
+                    >reverse_append >reverse_append >reverse_cons
+                    >reverse_reverse >reverse_reverse >reverse_reverse
+                    >(?:(la@[〈c',false〉])@((((lb@[〈grid,false〉])@l2)@la)@[〈d',false〉])@rs3
+                       =((la@〈c',false〉::lb)@([〈grid,false〉]@l2@la@[〈d',false〉]@rs3)))
                     [|>associative_append >associative_append 
                       >associative_append >associative_append >associative_append
-                      >associative_append >associative_append % ]
-                    <H2 normalize (* <Hrs3 *)
-                    >associative_append >associative_append >associative_append
-                    @eq_f normalize @eq_f >associative_append
-                    >associative_append @eq_f normalize @eq_f
-                    >(append_cons ? 〈d',false〉) >associative_append
-                    <Heq1 >Hl4 <associative_append <append_cons
-                    <H3
-                    >associative_append normalize
-                    >associative_append normalize %
+                      >associative_append % ]
+                    <H2 normalize in ⊢ (??%?); >Hrs3
+                    >associative_append >associative_append normalize
+                    >associative_append >associative_append
+                    @eq_f @eq_f @eq_f
+                    >(?:la@(〈d',false〉::lc@〈comma,false〉::rs32)@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs0 = 
+                        (la@〈d',false〉::lc)@〈comma,false〉::rs32@〈bar,false〉::〈d2,true〉::rs3'@〈grid,false〉::rs0 )
+                    [| >associative_append normalize >associative_append % ]
+                    <H3 %
                    ]
                  ]
               ]
-           ]
-         ]
        |* #Hnobars #Htapee >Htapee -Htapee *
          [whd in ⊢ (%→?); * #tapef * whd in ⊢ (%→?); #Htapef
           cases (Htapef … (refl …)) -Htapef #_ #Htapef >Htapef -Htapef
-          whd in ⊢ (%→?); #Htapeout %2
-          >(Htapeout … (refl …)) %
-           [ % 
-             [ @Hnoteq 
-             | whd #x #Hx @Hnobars @memb_append_l2 @memb_cons //
-             ]
-           | %
-           ] 
+          whd in ⊢ (%→?); #Htapeout %2 % 
+          [% [//] whd #x #Hx @Hnobars @memb_append_l2 @memb_cons //
+          | >(Htapeout … (refl …)) % ]
          |whd in ⊢ (%→?); * #tapef * whd in ⊢ (%→?); #Htapef
           cases (Htapef … (refl …)) -Htapef 
           whd in ⊢ ((??%?)→?); #Htemp destruct (Htemp) 
          ]
        |(* no marks in table *)
         #x #membx @(no_marks_in_table … Htable) 
-        @memb_append_l2 @memb_cons 
+        @memb_append_l2
         cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut >Hcut
         >H3 >associative_append @memb_append_l2 @memb_cons @membx
        |(* no grids in table *)
         #x #membx @(no_grids_in_table … Htable) 
-        @memb_append_l2 @memb_cons 
+        @memb_append_l2
         cut (∀A,l1,l2.∀a:A. a::l1@l2=(a::l1)@l2) [//] #Hcut >Hcut
         >H3 >associative_append @memb_append_l2 @memb_cons @membx
        |whd in ⊢ (??%?); >(bit_or_null_not_grid … Hd') >(bit_or_null_not_bar … Hd') %
        ]
      ]
    |#x #membx @(no_marks_in_table … Htable) 
-    @memb_append_l2 @memb_cons @memb_cons @memb_append_l1 @membx 
+    @memb_append_l2 @memb_cons @memb_append_l1 @membx 
    |#x #membx @(no_marks_in_table … Htable) 
-    cases (memb_append … membx) -membx #membx
-     [@memb_append_l1 @membx | @memb_append_l2 >(memb_single … membx) @memb_hd]
-   |>associative_append %
+    @memb_append_l1 @membx
+   |%
    ]
  ]
 qed.
 
-
 (* 
   MATCH TUPLE
 
@@ -763,23 +830,190 @@ qed.
   current configuration is found
 *)
 
-definition match_tuple ≝  whileTM ? match_tuple_step (inr … (inl … (inr … 0))).
+definition match_tuple ≝  whileTM ? match_tuple_step (inr … (inl … (inr … start_nop))).
+
+lemma is_grid_true : ∀c.is_grid c = true → c = grid.
+* normalize [ #b ] #H // destruct (H)
+qed.
+
+(* possible variante ? 
+definition weakR_match_tuple ≝ λt1,t2.
+  (∀ls,cur,rs,b. t1 = midtape STape ls 〈grid,b〉 rs → t2 = t1) ∧
+  (∀c,l1,c1,l2,l3,ls0,rs0,n.
+  t1 = midtape STape (〈grid,false〉::ls0) 〈bit c,true〉 rs 
+    (l1@〈grid,false〉::l2@〈bit c1,true〉::l3@〈grid,false〉::rs0) → 
+  only_bits_or_nulls l1 → no_marks l1 → S n = |l1| →
+  table_TM (S n) (l2@〈c1,false〉::l3) → 
+  (* facciamo match *)
+  (∃l4,newc,mv,l5.
+   〈c1,false〉::l3 = l4@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5 ∧
+   t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
+        (l2@l4@〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l5@
+        〈grid,false〉::rs0))
+  ∨
+  (* non facciamo match su nessuna tupla;
+     non specifichiamo condizioni sul nastro di output, perché
+     non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
+  (current ? t2 = Some ? 〈grid,true〉 ∧
+   ∀l4,newc,mv,l5.
+   〈c1,false〉::l3 ≠ l4@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5)).  
+*) 
+
+definition R_match_tuple0 ≝ λt1,t2.
+  ∀ls,cur,rs.
+  t1 = midtape STape ls cur rs → 
+  (is_grid (\fst cur) = true → t2 = t1) ∧
+  (∀c,l1,c1,l2,l3,ls0,rs0,n.
+  ls = 〈grid,false〉::ls0 → 
+  cur = 〈c,true〉 → 
+  rs = l1@〈grid,false〉::l2@〈bar,false〉::〈c1,true〉::l3@〈grid,false〉::rs0 → 
+  is_bit c = true → is_bit c1 = true → 
+  only_bits_or_nulls l1 → no_marks l1 → S n = |l1| →
+  table_TM (S n) (l2@〈bar,false〉::〈c1,false〉::l3) → 
+  (* facciamo match *)
+  (∃l4,newc,mv,l5.
+   〈bar,false〉::〈c1,false〉::l3 = l4@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5 ∧
+   t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls0) 〈grid,false〉
+        (l2@l4@〈bar,false〉::〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l5@
+        〈grid,false〉::rs0))
+  ∨
+  (* non facciamo match su nessuna tupla;
+     non specifichiamo condizioni sul nastro di output, perché
+     non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
+  (current ? t2 = Some ? 〈grid,true〉 ∧
+   ∀l4,newc,mv,l5.
+   〈bar,false〉::〈c1,false〉::l3 ≠ l4@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l5)).  
+
+axiom table_bit_after_bar : 
+  ∀n,l1,c,l2.table_TM n (l1@〈bar,false〉::〈c,false〉::l2) → is_bit c = true.
+
+lemma wsem_match_tuple : WRealize ? match_tuple R_match_tuple0.
+#intape #k #outc #Hloop 
+lapply (sem_while … sem_match_tuple_step intape k outc Hloop) [%] -Hloop
+* #ta * #Hstar @(star_ind_l ??????? Hstar)
+[ #tb whd in ⊢ (%→?); #Hleft
+  #ls #cur #rs #Htb cases (Hleft … Htb) #Hgrid #Houtc %
+  [ #_ @Houtc 
+  | #c #l1 #c1 #l2 #l3 #ls0 #rs0 #n #Hls #Hcur #Hrs 
+    >Hcur in Hgrid; #Hgrid >(is_grid_true … Hgrid) normalize in ⊢ (%→?);
+    #Hc destruct (Hc)
+  ]
+| (* in the interesting case, we execute a true iteration, then we restart the
+     while cycle, finally we end with a false iteration *)
+  #tb #tc #td whd in ⊢ (%→?); #Htc
+  #Hstar1 #IH whd in ⊢ (%→?); #Hright lapply (IH Hright) -IH whd in ⊢ (%→?); #IH
+  #ls #cur #rs #Htb %
+  [ (* cur can't be true because we assume at least one iteration *)
+    #Hcur cases (Htc … Htb) * #Hfalse @False_ind @Hfalse @(is_grid_true … Hcur)
+  | (* current and a tuple are marked *)
+   #c #l1 #c1 #l2 #l3 #ls0 #rs0 #n #Hls #Hcur #Hrs #Hc #Hc1 #Hl1bitnull #Hl1marks 
+   #Hl1len #Htable cases (Htc … Htb) -Htc -Htb * #_ #Htc
+   (* expose the marked tuple in table *)
+   cut (∃la,lb,mv,lc.l3 = la@〈comma,false〉::lb@〈comma,false〉::mv::lc ∧
+         S n = |la| ∧ only_bits_or_nulls la)
+   [@daemon] * #la * #lb * #mv * #lc * * #Hl3 #Hlalen #Hlabitnull
+   >Hl3 in Htable; >append_cons #Htable
+   >(?: l2@〈bar,false〉::〈c1,true〉::l3@〈grid,false〉::rs0
+      = (l2@[〈bar,false〉])@〈c1,true〉::la@〈comma,false〉::(lb@〈comma,false〉::mv::
+         lc)@〈grid,false〉::rs0) in Hrs;
+   [| >associative_append normalize >Hl3
+      >associative_append normalize % ] #Hrs
+   cases (Htc ????????? Hl1bitnull Hl1marks ?? Hlabitnull Hl1len ? Htable Hls Hcur Hrs)
+   [5: <Hl1len @Hlalen
+   |4: whd in ⊢ (??%?); >Hc1 %
+   |3: whd in ⊢ (??%?); >Hc %
+   |-Htc *
+     [ (* case 1: match successful *)
+       * #Heq #Htc % %{[]} %{lb} %{mv} %{lc} destruct (Heq) %
+       [%
+       | cases (IH … Htc) -IH #Houtc #_ >(Houtc (refl ??)) 
+         >Htc @eq_f normalize >associative_append normalize
+         >associative_append normalize %
+       ]     
+     | (* case 2: tuples don't match, we still have other tuples to try *)
+       * #Hdiff * #c2 * #l5 * #l6 * #Heqlblc #Htc
+       cases (IH ??? … Htc) -IH #_ #IH 
+       (* by induction hypothesis *)
+       lapply (IH ? l1 c2 (l2@〈bar,false〉::〈c1,false〉::la@〈comma,false〉::l5) l6 ? rs0 n (refl ??) (refl ??) ???????)
+       [ generalize in match Htable;
+         >associative_append normalize 
+         >associative_append normalize >Heqlblc
+         >associative_append normalize //
+       | @Hl1len
+       | @Hl1marks
+       | @Hl1bitnull
+       | (*???*) @daemon
+       | @Hc
+       | >associative_append normalize 
+         >associative_append normalize
+         >associative_append %
+       |-IH * 
+         [ (* the while finally matches a tuple *)
+           * #l7 * #newc * #mv0 * #l8 * #Hl7l8 #Houtc %
+           >Heqlblc @(ex_intro ?? (〈bar,false〉::〈c1,false〉::la@〈comma,false〉::l5@l7))
+           %{newc} %{mv0} %{l8} %
+           [ normalize >Hl7l8 >associative_append normalize 
+             >associative_append %
+           | >Houtc @eq_f >associative_append normalize
+             >associative_append normalize >associative_append 
+             normalize >associative_append %
+           ]
+         | (* the while fails finding a tuple: there are no matches in the whole table *)
+           * #Houtc #Hdiff1 %2 %
+           [ @Houtc
+           | #l50 #newc #mv0 #l51 >Heqlblc 
+             @daemon
+           ]
+         ]
+       ]
+     ]
+   | (* match failed and there is no next tuple: the next while cycle will just exit *)
+     * * #Hdiff #Hnobars generalize in match (refl ? tc);
+     cases tc in ⊢ (???% → %);
+     [ #_ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
+     |2,3: #x #xs #_ normalize in ⊢ (??%?→?); #Hfalse destruct (Hfalse) ]
+     #ls1 #cur1 #rs1 #Htc normalize in ⊢ (??%?→?); #Hcur1
+     cases (IH … Htc) -IH #IH #_ %2 %
+     [ destruct (Hcur1) >IH [ >Htc % | % ]
+     | #l4 #newc #mv0 #l5
+       (* no_bars except the first one, where the tuple does not match ⇒ 
+          no match *)
+        @daemon
+     ]
+   ]
+ ]
+qed.
 
 definition R_match_tuple ≝ λt1,t2.
   ∀ls,c,l1,c1,l2,rs,n.
-  is_bit c = true → only_bits_or_nulls l1 → is_bit c1 = true → n = |l1| →
-  table_TM (S n) (〈c1,false〉::l2) → 
+  is_bit c = true → is_bit c1 = true →
+  only_bits_or_nulls l1 → no_marks l1 → S n = |l1| → 
+  table_TM (S n) (〈bar,false〉::〈c1,false〉::l2) → 
   t1 = midtape STape (〈grid,false〉::ls) 〈c,true〉 
-         (l1@〈grid,false〉::〈c1,true〉::l2@〈grid,false〉::rs) → 
+         (l1@〈grid,false〉::〈bar,false〉::〈c1,true〉::l2@〈grid,false〉::rs) → 
   (* facciamo match *)
   (∃l3,newc,mv,l4.
-   〈c1,false〉::l2 = l3@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4 ∧
+   〈bar,false〉::〈c1,false〉::l2 = l3@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4 ∧
    t2 = midtape ? (reverse ? l1@〈c,false〉::〈grid,false〉::ls) 〈grid,false〉
-        (l3@〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l4@〈grid,false〉::rs))
+        (l3@〈bar,false〉::〈c,false〉::l1@〈comma,true〉::newc@〈comma,false〉::mv::l4@〈grid,false〉::rs))
   ∨
   (* non facciamo match su nessuna tupla;
      non specifichiamo condizioni sul nastro di output, perché
      non eseguiremo altre operazioni, quindi il suo formato non ci interessa *)
   (current ? t2 = Some ? 〈grid,true〉 ∧
    ∀l3,newc,mv,l4.
-   〈c1,false〉::l2 ≠ l3@〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4).  
+   〈bar,false〉::〈c1,false〉::l2 ≠ l3@〈bar,false〉::〈c,false〉::l1@〈comma,false〉::newc@〈comma,false〉::mv::l4). 
+
+(* we still haven't proved termination *)
+axiom sem_match_tuple0 : Realize ? match_tuple R_match_tuple0.
+
+axiom Realize_to_Realize : 
+  ∀alpha,M,R1,R2.(∀t1,t2.R1 t1 t2 → R2 t1 t2) → Realize alpha M R1 → Realize alpha M R2.
+
+lemma sem_match_tuple : Realize ? match_tuple R_match_tuple.
+generalize in match sem_match_tuple0; @Realize_to_Realize
+#t1 #t2 #HR #ls #c #l1 #c1 #l2 #rs #n #Hc #Hc1 #Hl1bitsnulls #Hl1marks #Hl1len #Htable #Ht1
+cases (HR … Ht1) -HR #_ #HR
+@(HR ??? [] … (refl ??) (refl ??) (refl ??) Hc Hc1 Hl1bitsnulls Hl1marks
+          Hl1len  Htable)
+qed.
\ No newline at end of file