]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/turing/universal/marks.ma
porting of move_char_c
[helm.git] / matita / matita / lib / turing / universal / marks.ma
index bfb0cf4bd885fb2b797ce5eb945a42a6a1dd26d3..70000e3996690e2837ead12ca3c932cfa581ae12 100644 (file)
@@ -16,6 +16,7 @@
 
 include "turing/while_machine.ma".
 include "turing/if_machine.ma".
+include "turing/universal/alphabet.ma".
 include "turing/universal/tests.ma".
 
 (* ADVANCE TO MARK (right)
@@ -29,17 +30,21 @@ include "turing/universal/tests.ma".
 
 definition atm_states ≝ initN 3.
 
+definition atm0 : atm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
+definition atm1 : atm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
+definition atm2 : atm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
+
 definition atmr_step ≝ 
   λalpha:FinSet.λtest:alpha→bool.
   mk_TM alpha atm_states
   (λp.let 〈q,a〉 ≝ p in
    match a with
-   [ None ⇒ 〈1, None ?〉
+   [ None ⇒ 〈atm1, None ?〉
    | Some a' ⇒ 
      match test a' with
-     [ true ⇒ 〈1,None ?〉
-     | false ⇒ 〈2,Some ? 〈a',R〉〉 ]])
-  O (λx.notb (x == 0)).
+     [ true ⇒ 〈atm1,None ?〉
+     | false ⇒ 〈atm2,Some ? 〈a',R〉〉 ]])
+  atm0 (λx.notb (x == atm0)).
 
 definition Ratmr_step_true ≝ 
   λalpha,test,t1,t2.
@@ -56,43 +61,45 @@ definition Ratmr_step_false ≝
 lemma atmr_q0_q1 :
   ∀alpha,test,ls,a0,rs. test a0 = true → 
   step alpha (atmr_step alpha test)
-    (mk_config ?? 0 (midtape … ls a0 rs)) =
-  mk_config alpha (states ? (atmr_step alpha test)) 1
+    (mk_config ?? atm0 (midtape … ls a0 rs)) =
+  mk_config alpha (states ? (atmr_step alpha test)) atm1
     (midtape … ls a0 rs).
-#alpha #test #ls #a0 #ts #Htest normalize >Htest %
+#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?);
+whd in match (trans … 〈?,?〉); >Htest %
 qed.
      
 lemma atmr_q0_q2 :
   ∀alpha,test,ls,a0,rs. test a0 = false → 
   step alpha (atmr_step alpha test)
-    (mk_config ?? 0 (midtape … ls a0 rs)) =
-  mk_config alpha (states ? (atmr_step alpha test)) 2
+    (mk_config ?? atm0 (midtape … ls a0 rs)) =
+  mk_config alpha (states ? (atmr_step alpha test)) atm2
     (mk_tape … (a0::ls) (option_hd ? rs) (tail ? rs)).
-#alpha #test #ls #a0 #ts #Htest normalize >Htest cases ts //
+#alpha #test #ls #a0 #ts #Htest whd in ⊢ (??%?);
+whd in match (trans … 〈?,?〉); >Htest cases ts //
 qed.
 
 lemma sem_atmr_step :
   ∀alpha,test.
   accRealize alpha (atmr_step alpha test) 
-    2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test).
+    atm2 (Ratmr_step_true alpha test) (Ratmr_step_false alpha test).
 #alpha #test *
 [ @(ex_intro ?? 2)
-  @(ex_intro ?? (mk_config ?? 1 (niltape ?))) %
-  [ % // #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (leftof ? a al)))
-  % [ % // #Hfalse destruct | #_ % // % % ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (rightof ? a al)))
-  % [ % // #Hfalse destruct | #_ % // % % ]
+  @(ex_intro ?? (mk_config ?? atm1 (niltape ?))) %
+  [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
+| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (leftof ? a al)))
+  % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
+| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? atm1 (rightof ? a al)))
+  % [ % // whd in ⊢ ((??%%)→?); #Hfalse destruct | #_ % // % % ]
 | #ls #c #rs @(ex_intro ?? 2)
   cases (true_or_false (test c)) #Htest
-  [ @(ex_intro ?? (mk_config ?? 1 ?))
+  [ @(ex_intro ?? (mk_config ?? atm1 ?))
     [| % 
       [ % 
         [ whd in ⊢ (??%?); >atmr_q0_q1 //
-        | #Hfalse destruct ]
+        | whd in ⊢ ((??%%)→?); #Hfalse destruct ]
       | #_ % // %2 @(ex_intro ?? c) % // ]
     ]
-  | @(ex_intro ?? (mk_config ?? 2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs))))
+  | @(ex_intro ?? (mk_config ?? atm2 (mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs))))
     % 
     [ %
       [ whd in ⊢ (??%?); >atmr_q0_q2 //
@@ -115,7 +122,7 @@ definition R_adv_to_mark_r ≝ λalpha,test,t1,t2.
      t2 = midtape ? (reverse ? rs1@c::ls) b rs2))).
      
 definition adv_to_mark_r ≝ 
-  λalpha,test.whileTM alpha (atmr_step alpha test) 2.
+  λalpha,test.whileTM alpha (atmr_step alpha test) atm2.
 
 lemma wsem_adv_to_mark_r :
   ∀alpha,test.
@@ -187,15 +194,18 @@ qed.
  
 definition mark_states ≝ initN 2.
 
+definition ms0 : mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
+definition ms1 : mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
+
 definition mark ≝ 
   λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mark_states
   (λp.let 〈q,a〉 ≝ p in
     match a with
-    [ None ⇒ 〈1,None ?〉
-    | Some a' ⇒ match q with
-      [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',true〉,N〉〉
-      | S q ⇒ 〈1,None ?〉 ] ])
-  O (λq.q == 1).
+    [ None ⇒ 〈ms1,None ?〉
+    | Some a' ⇒ match (pi1 … q) with
+      [ O ⇒ let 〈a'',b〉 ≝ a' in 〈ms1,Some ? 〈〈a'',true〉,N〉〉
+      | S q ⇒ 〈ms1,None ?〉 ] ])
+  ms0 (λq.q == ms1).
   
 definition R_mark ≝ λalpha,t1,t2.
   ∀ls,c,b,rs.
@@ -222,16 +232,18 @@ qed.
 *)
 
 definition move_states ≝ initN 2.
+definition move0 : move_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)).
+definition move1 : move_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)).
 
 definition move_r ≝ 
   λalpha:FinSet.mk_TM alpha move_states
   (λp.let 〈q,a〉 ≝ p in
     match a with
-    [ None ⇒ 〈1,None ?〉
-    | Some a' ⇒ match q with
-      [ O ⇒ 〈1,Some ? 〈a',R〉〉
-      | S q ⇒ 〈1,None ?〉 ] ])
-  O (λq.q == 1).
+    [ None ⇒ 〈move1,None ?〉
+    | Some a' ⇒ match (pi1 … q) with
+      [ O ⇒ 〈move1,Some ? 〈a',R〉〉
+      | S q ⇒ 〈move1,None ?〉 ] ])
+  move0 (λq.q == move1).
   
 definition R_move_r ≝ λalpha,t1,t2.
   ∀ls,c,rs.
@@ -262,11 +274,11 @@ definition move_l ≝
   λalpha:FinSet.mk_TM alpha move_states
   (λp.let 〈q,a〉 ≝ p in
     match a with
-    [ None ⇒ 〈1,None ?〉
-    | Some a' ⇒ match q with
-      [ O ⇒ 〈1,Some ? 〈a',L〉〉
-      | S q ⇒ 〈1,None ?〉 ] ])
-  O (λq.q == 1).
+    [ None ⇒ 〈move1,None ?〉
+    | Some a' ⇒ match pi1 … q with
+      [ O ⇒ 〈move1,Some ? 〈a',L〉〉
+      | S q ⇒ 〈move1,None ?〉 ] ])
+  move0 (λq.q == move1).
   
 definition R_move_l ≝ λalpha,t1,t2.
   ∀ls,c,rs.
@@ -296,18 +308,22 @@ qed.
  
 definition mrm_states ≝ initN 3.
 
+definition mrm0 : mrm_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
+definition mrm1 : mrm_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
+definition mrm2 : mrm_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
+
 definition move_right_and_mark ≝ 
   λalpha:FinSet.mk_TM (FinProd … alpha FinBool) mrm_states
   (λp.let 〈q,a〉 ≝ p in
     match a with
-    [ None ⇒ 〈2,None ?〉
-    | Some a' ⇒ match q with
-      [ O ⇒ 〈1,Some ? 〈a',R〉〉
+    [ None ⇒ 〈mrm2,None ?〉
+    | Some a' ⇒ match pi1 … q with
+      [ O ⇒ 〈mrm1,Some ? 〈a',R〉〉
       | S q ⇒ match q with
         [ O ⇒ let 〈a'',b〉 ≝ a' in
-              〈2,Some ? 〈〈a'',true〉,N〉〉
-        | S _ ⇒ 〈2,None ?〉 ] ] ])
-  O (λq.q == 2).
+              〈mrm2,Some ? 〈〈a'',true〉,N〉〉
+        | S _ ⇒ 〈mrm2,None ?〉 ] ] ])
+  mrm0 (λq.q == mrm2).
   
 definition R_move_right_and_mark ≝ λalpha,t1,t2.
   ∀ls,c,d,b,rs.
@@ -336,15 +352,19 @@ qed.
  
 definition clear_mark_states ≝ initN 3.
 
+definition clear0 : clear_mark_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
+definition clear1 : clear_mark_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
+definition claer2 : clear_mark_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
+
 definition clear_mark ≝ 
   λalpha:FinSet.mk_TM (FinProd … alpha FinBool) clear_mark_states
   (λp.let 〈q,a〉 ≝ p in
     match a with
-    [ None ⇒ 〈1,None ?〉
-    | Some a' ⇒ match q with
-      [ O ⇒ let 〈a'',b〉 ≝ a' in 〈1,Some ? 〈〈a'',false〉,N〉〉
-      | S q ⇒ 〈1,None ?〉 ] ])
-  O (λq.q == 1).
+    [ None ⇒ 〈clear1,None ?〉
+    | Some a' ⇒ match pi1 … q with
+      [ O ⇒ let 〈a'',b〉 ≝ a' in 〈clear1,Some ? 〈〈a'',false〉,N〉〉
+      | S q ⇒ 〈clear1,None ?〉 ] ])
+  clear0 (λq.q == clear1).
   
 definition R_clear_mark ≝ λalpha,t1,t2.
   ∀ls,c,b,rs.
@@ -542,30 +562,6 @@ cases (sem_seq ????? (sem_adv_mark_r …)
     | >reverse_append #Htc @Htc ]
   ]
 qed.
-  
-inductive unialpha : Type[0] ≝ 
-| bit : bool → unialpha
-| comma : unialpha
-| bar : unialpha
-| grid : unialpha.
-
-definition unialpha_eq ≝ 
-  λa1,a2.match a1 with
-  [ bit x ⇒ match a2 with [ bit y ⇒ ¬ xorb x y | _ ⇒ false ]
-  | comma ⇒ match a2 with [ comma ⇒ true | _ ⇒ false ]
-  | bar ⇒ match a2 with [ bar ⇒ true | _ ⇒ false ]
-  | grid ⇒ match a2 with [ grid ⇒ true | _ ⇒ false ] ].
-  
-definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?.
-* [ #x * [ #y cases x cases y normalize % // #Hfalse destruct
-         | *: normalize % #Hfalse destruct ]
-  |*: * [1,5,9,13: #y ] normalize % #H1 destruct % ]
-qed.
-
-definition FSUnialpha ≝ 
-  mk_FinSet DeqUnialpha [bit true;bit false;comma;bar;grid] ?.
-@daemon
-qed.
 
 (* 
    MATCH AND ADVANCE(f)
@@ -701,17 +697,16 @@ definition comp_step ≝
   ifTM ? (test_char ? (is_marked ?))
   (single_finalTM … (comp_step_subcase FSUnialpha 〈bit false,true〉
     (comp_step_subcase FSUnialpha 〈bit true,true〉
-      (clear_mark …))))
+      (comp_step_subcase FSUnialpha 〈null,true〉
+        (clear_mark …)))))
   (nop ?)
   tc_true.
   
-definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ].
-  
 definition R_comp_step_true ≝ 
   λt1,t2.
     ∀l0,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) l0 c rs → 
     ∃c'. c = 〈c',true〉 ∧
-    ((is_bit c' = true ∧
+    ((bit_or_null c' = true ∧
      ∀a,l1,c0,a0,l2.
       rs = 〈a,false〉::l1@〈c0,true〉::〈a0,false〉::l2 → 
       (∀c.memb ? c l1 = true → is_marked ? c = false) → 
@@ -720,7 +715,7 @@ definition R_comp_step_true ≝
       (c0 ≠ c' ∧
        t2 = midtape (FinProd … FSUnialpha FinBool) 
         (reverse ? l1@〈a,false〉::〈c',true〉::l0) 〈c0,false〉 (〈a0,false〉::l2))) ∨
-     (is_bit c' = false ∧ t2 = midtape ? l0 〈c',false〉 rs)).
+     (bit_or_null c' = false ∧ t2 = midtape ? l0 〈c',false〉 rs)).
 
 definition R_comp_step_false ≝ 
   λt1,t2.
@@ -728,13 +723,14 @@ definition R_comp_step_false ≝
    is_marked ? c = false ∧ t2 = t1.
    
 lemma sem_comp_step : 
-  accRealize ? comp_step (inr … (inl … (inr … 0))) 
+  accRealize ? comp_step (inr … (inl … (inr … start_nop))) 
     R_comp_step_true R_comp_step_false.
 #intape
 cases (acc_sem_if … (sem_test_char ? (is_marked ?))
         (sem_comp_step_subcase FSUnialpha 〈bit false,true〉 ??
           (sem_comp_step_subcase FSUnialpha 〈bit true,true〉 ?? 
-            (sem_clear_mark …)))
+            (sem_comp_step_subcase FSUnialpha 〈null,true〉 ?? 
+              (sem_clear_mark …))))
         (sem_nop …) intape)
 #k * #outc * * #Hloop #H1 #H2
 @(ex_intro ?? k) @(ex_intro ?? outc) %
@@ -762,13 +758,24 @@ cases (acc_sem_if … (sem_test_char ? (is_marked ?))
         [ @sym_not_eq //
         | @Houtc ]
       ]
-    | * #Hc' whd in ⊢ (%→?); #Helse2 %2 %
-      [ generalize in match Hc'; generalize in match Hc;
-        cases c'
-        [ * [ #_ #Hfalse @False_ind @(absurd ?? Hfalse) %
-            | #Hfalse @False_ind @(absurd ?? Hfalse) % ]
-        |*: #_ #_ % ]
-      | @(Helse2 … Hta)
+    | * #Hc' #Helse2 cases (Helse2 … Hta)
+      [ * #Hc'' #H1 % % [destruct (Hc'') % ]
+        #a #l1 #c0 #a0 #l2 #Hrs >Hrs in Hintape; #Hintape #Hl1
+        cases (H1 … Hl1 Hrs)
+        [ * #Htmp >Htmp -Htmp #Houtc % % // @Houtc
+        | * #Hneq #Houtc %2 %
+          [ @sym_not_eq //
+          | @Houtc ]
+        ]
+      | * #Hc'' whd in ⊢ (%→?); #Helse3 %2 %
+        [ generalize in match Hc''; generalize in match Hc'; generalize in match Hc;
+          cases c'
+          [ * [ #_ #Hfalse @False_ind @(absurd ?? Hfalse) %
+              | #Hfalse @False_ind @(absurd ?? Hfalse) % ]
+          | #_ #_ #Hfalse @False_ind @(absurd ?? Hfalse) %
+          |*: #_ #_ #_ % ]
+        | @(Helse3 … Hta)
+        ]
       ]
     ]
   ]
@@ -781,7 +788,7 @@ cases (acc_sem_if … (sem_test_char ? (is_marked ?))
 qed.
 
 definition compare ≝ 
-  whileTM ? comp_step (inr … (inl … (inr … 0))).
+  whileTM ? comp_step (inr … (inl … (inr … start_nop))).
 
 (*
 definition R_compare :=
@@ -832,16 +839,16 @@ RIFIUTO: c ≠ d
 definition R_compare :=
   λt1,t2.
   ∀ls,c,rs.t1 = midtape (FinProd … FSUnialpha FinBool) ls c rs → 
-  (∀c'.is_bit c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧
+  (∀c'.bit_or_null c' = false → c = 〈c',true〉 → t2 = midtape ? ls 〈c',false〉 rs) ∧
   (∀c'. c = 〈c',false〉 → t2 = t1) ∧
   ∀b,b0,bs,b0s,l1,l2.
   |bs| = |b0s| → 
-  (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → is_bit (\fst c) = true) → 
-  (∀c.memb (FinProd … FSUnialpha FinBool) c b0s = true → is_bit (\fst c) = true) → 
+  (∀c.memb (FinProd … FSUnialpha FinBool) c bs = true → bit_or_null (\fst c) = true) → 
+  (∀c.memb (FinProd … FSUnialpha FinBool) c b0s = true → bit_or_null (\fst c) = true) → 
   (∀c.memb ? c bs = true → is_marked ? c = false) → 
   (∀c.memb ? c b0s = true → is_marked ? c = false) → 
   (∀c.memb ? c l1 = true → is_marked ? c = false) → 
-  c = 〈b,true〉 → is_bit b = true → 
+  c = 〈b,true〉 → bit_or_null b = true → 
   rs = bs@〈grid,false〉::l1@〈b0,true〉::b0s@〈comma,false〉::l2 → 
   (〈b,true〉::bs = 〈b0,true〉::b0s ∧
    t2 = midtape ? (reverse ? bs@〈b,false〉::ls)
@@ -898,7 +905,7 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%]
   ]
 | #tapea #tapeb #tapec #Hleft #Hright #IH #Htapec lapply (IH Htapec) -Htapec -IH #IH
   whd in Hleft; #ls #c #rs #Htapea cases (Hleft … Htapea) -Hleft
-  #c' * #Hc >Hc cases (true_or_false (is_bit c')) #Hc'
+  #c' * #Hc >Hc cases (true_or_false (bit_or_null c')) #Hc'
   [2: * 
     [ * >Hc' #H @False_ind destruct (H)
     | * #_ #Htapeb cases (IH … Htapeb) * #_ #H #_ %
@@ -921,147 +928,79 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%]
        @(list_cases_2 … Hlen)
        [ #Hbs #Hb0s generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?);
        -Hrs #Hrs normalize in Hrs; #Hleft cases (Hleft ????? Hrs ?) -Hleft
-         [ * #Heqb #Htapeb cases (IH … Htapeb) -IH * #IH #_ #IH1
+         [ * #Heqb #Htapeb cases (IH … Htapeb) -IH * #IH #_ #_
           % %
             [ >Heqb >Hbs >Hb0s %
             | >Hbs >Hb0s @IH %
             ] 
-          |* #Hneqb #Htapeb %2
-           @(ex_intro … [ ]) @(ex_intro … b)
-           @(ex_intro … b0) @(ex_intro … [ ]) 
-           @(ex_intro … [ ]) %
+         |* #Hneqb #Htapeb %2
+          @(ex_intro … [ ]) @(ex_intro … b)
+          @(ex_intro … b0) @(ex_intro … [ ]) 
+          @(ex_intro … [ ]) %
             [ % [ % [@sym_not_eq //| >Hbs %] | >Hb0s %]
-            |
-          #l3 #c0 #Hyp >Hbs >Hb0s 
-         cases (IH b b0 bs l1 l2 Hlen ?????
-     
-     >(\P Hc') whd in ⊢ (??%?→?); #Hfalse destruct (Hfalse)
-    ]
-qed.
-                   
-
-
-(*
-l0 x* a l1 x0* a0 l2 ------> l0 x a* l1 x0 a0* l2
-   ^                               ^
-
-if current (* x *) = #
-   then 
-   else if x = 0
-      then move_right; ----
-           adv_to_mark_r;
-           if current (* x0 *) = 0
-              then advance_mark ----
-                   adv_to_mark_l;
-                   advance_mark
-              else STOP
-      else x = 1 (* analogo *)
-
-*)
-
-
-(*
-   MARK NEXT TUPLE machine
-   (partially axiomatized)
-   
-   marks the first character after the first bar (rightwards)
- *)
-
-axiom myalpha : FinSet.
-axiom is_bar : FinProd … myalpha FinBool → bool.
-axiom is_grid : FinProd … myalpha FinBool → bool.
-definition bar_or_grid ≝ λc.is_bar c ∨ is_grid c.
-axiom bar : FinProd … myalpha FinBool.
-axiom grid : FinProd … myalpha FinBool.
-
-definition mark_next_tuple ≝ 
-  seq ? (adv_to_mark_r ? bar_or_grid)
-     (ifTM ? (test_char ? is_bar)
-       (move_r_and_mark ?) (nop ?) 1).
-
-definition R_mark_next_tuple ≝ 
-  λt1,t2.
-    ∀ls,c,rs1,rs2.
-    (* c non può essere un separatore ... speriamo *)
-    t1 = midtape ? ls c (rs1@grid::rs2) → 
-    memb ? grid rs1 = false → bar_or_grid c = false → 
-    (∃rs3,rs4,d,b.rs1 = rs3 @ bar :: rs4 ∧
-      memb ? bar rs3 = false ∧ 
-      Some ? 〈d,b〉 = option_hd ? (rs4@grid::rs2) ∧
-      t2 = midtape ? (bar::reverse ? rs3@c::ls) 〈d,true〉 (tail ? (rs4@grid::rs2)))
-    ∨
-    (memb ? bar rs1 = false ∧ 
-     t2 = midtape ? (reverse ? rs1@c::ls) grid rs2).
-     
-axiom tech_split :
-  ∀A:DeqSet.∀f,l.
-   (∀x.memb A x l = true → f x = false) ∨
-   (∃l1,c,l2.f c = true ∧ l = l1@c::l2 ∧ ∀x.memb ? x l1 = true → f c = false).
-(*#A #f #l elim l
-[ % #x normalize #Hfalse *)
-     
-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 ? is_bar) (mark ?) (nop ?) 1) ????)
-[@sem_if //
-| //
-|||#Hif cases (Hif intape) -Hif
-   #j * #outc * #Hloop * #ta * #Hleft #Hright
-   @(ex_intro ?? j) @ex_intro [|% [@Hloop] ]
-   -Hloop
-   #ls #c #rs1 #rs2 #Hrs #Hrs1 #Hc
-   cases (Hleft … Hrs)
-   [ * #Hfalse >Hfalse in Hc; #Htf destruct (Htf)
-   | * #_ #Hta cases (tech_split ? is_bar rs1)
-     [ #H1 lapply (Hta rs1 grid rs2 (refl ??) ? ?)
-       [ (* Hrs1, H1 *) @daemon
-       | (* bar_or_grid grid = true *) @daemon
-       | -Hta #Hta cases Hright
-         [ * #tb * whd in ⊢ (%→?); #Hcurrent
-           @False_ind cases(Hcurrent grid ?)
-           [ #Hfalse (* grid is not a bar *) @daemon
-           | >Hta % ]
-         | * #tb * whd in ⊢ (%→?); #Hcurrent
-           cases (Hcurrent grid ?)
-           [  #_ #Htb whd in ⊢ (%→?); #Houtc
-             %2 %
-             [ (* H1 *) @daemon
-             | >Houtc >Htb >Hta % ]
-           | >Hta % ]
+            | cases (IH … Htapeb) -IH * #_ #IH #_ >(IH ? (refl ??))
+              @Htapeb
+            ]
+         | @Hl1 ]
+      | * #b' #bitb' * #b0' #bitb0' #bs' #b0s' #Hbs #Hb0s 
+        generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?);
+        cut (bit_or_null b' = true ∧ bit_or_null b0' = true ∧ 
+             bitb' = false ∧ bitb0' = false)
+        [ % [ % [ % [ >Hbs in Hbs1; #Hbs1 @(Hbs1 〈b',bitb'〉) @memb_hd
+            | >Hb0s in Hb0s1; #Hb0s1 @(Hb0s1 〈b0',bitb0'〉) @memb_hd ]
+            | >Hbs in Hbs2; #Hbs2 @(Hbs2 〈b',bitb'〉) @memb_hd ]
+            | >Hb0s in Hb0s2; #Hb0s2 @(Hb0s2 〈b0',bitb0'〉) @memb_hd ]
+        | * * * #Ha #Hb #Hc #Hd >Hc >Hd
+          #Hrs #Hleft 
+          cases (Hleft b' (bs'@〈grid,false〉::l1) b0 b0' 
+                         (b0s'@〈comma,false〉::l2) ??) -Hleft
+          [ 3: >Hrs normalize @eq_f >associative_append %
+          | * #Hb0 #Htapeb cases (IH …Htapeb) -IH * #_ #_ #IH
+            cases (IH b' b0' bs' b0s' (l1@[〈b0,false〉]) l2 ??????? Ha ?) -IH
+            [ * #Heq #Houtc % %
+              [ >Hb0 @eq_f >Hbs in Heq; >Hb0s in ⊢ (%→?); #Heq
+                destruct (Heq) >Hb0s >Hc >Hd %
+              | >Houtc >Hbs >Hb0s >Hc >Hd >reverse_cons >associative_append
+                >associative_append %
+              ]
+            | * #la * #c' * #d' * #lb * #lc * * * #H1 #H2 #H3 #H4 %2
+              @(ex_intro … (〈b,false〉::la)) @(ex_intro … c') @(ex_intro … d')
+              @(ex_intro … lb) @(ex_intro … lc)
+              % [ % [ % // >Hbs >Hc >H2 % | >Hb0s >Hd >H3 >Hb0 % ] 
+                | >H4 >Hbs >Hb0s >Hc >Hd >Hb0 >reverse_append
+                  >reverse_cons >reverse_cons
+                  >associative_append >associative_append
+                  >associative_append >associative_append %
+                ]
+            | generalize in match Hlen; >Hbs >Hb0s
+              normalize #Hlen destruct (Hlen) @e0
+            | #c0 #Hc0 @Hbs1 >Hbs @memb_cons // 
+            | #c0 #Hc0 @Hb0s1 >Hb0s @memb_cons // 
+            | #c0 #Hc0 @Hbs2 >Hbs @memb_cons // 
+            | #c0 #Hc0 @Hb0s2 >Hb0s @memb_cons // 
+            | #c0 #Hc0 cases (memb_append … Hc0) 
+              [ @Hl1 | #Hc0' >(memb_single … Hc0') % ]
+            | %
+            | >associative_append >associative_append % ]
+         | * #Hneq #Htapeb %2
+            @(ex_intro … []) @(ex_intro … b) @(ex_intro … b0)
+            @(ex_intro … bs) @(ex_intro … b0s) %
+           [ % // % // @sym_not_eq // 
+           | >Hbs >Hb0s >Hc >Hd >reverse_cons >associative_append
+             >reverse_append in Htapeb; >reverse_cons
+             >associative_append >associative_append
+             #Htapeb <Htapeb
+             cases (IH … Htapeb) -Htapeb -IH * #_ #IH #_ @(IH ? (refl ??))
+           ]
+         | #c1 #Hc1 cases (memb_append … Hc1) #Hyp
+           [ @Hbs2 >Hbs @memb_cons @Hyp
+           | cases (orb_true_l … Hyp)
+             [ #Hyp2 >(\P Hyp2) %
+             | @Hl1
+             ]
+           ]
          ]
-       ]
-    | * #rs3 * #c0 * #rs4 * * #Hc0 #Hsplit #Hrs3
-      % @(ex_intro ?? rs3) @(ex_intro ?? rs4)
-     lapply (Hta rs3 c0 (rs4@grid::rs2) ???)
-     [ #x #Hrs3' (* Hrs1, Hrs3, Hsplit *) @daemon
-     | (* bar → bar_or_grid *) @daemon
-     | >Hsplit >associative_append % ] -Hta #Hta
-       cases Hright
-       [ * #tb * whd in ⊢ (%→?); #Hta'
-         whd in ⊢ (%→?); #Htb
-         cases (Hta' c0 ?)
-         [ #_ #Htb' >Htb' in Htb; #Htb
-           generalize in match Hsplit; -Hsplit
-           cases rs4 in Hta;
-           [ >(eq_pair_fst_snd … grid)
-             #Hta #Hsplit >(Htb … Hta)
-             >(?:c0 = bar)
-             [ @(ex_intro ?? (\fst grid)) @(ex_intro ?? (\snd grid))
-               % [ % [ % [ (* Hsplit *) @daemon |(*Hrs3*) @daemon ] | % ] | % ] 
-                     | (* Hc0 *) @daemon ]
-           | #r5 #rs5 >(eq_pair_fst_snd … r5)
-             #Hta #Hsplit >(Htb … Hta)
-             >(?:c0 = bar)
-             [ @(ex_intro ?? (\fst r5)) @(ex_intro ?? (\snd r5))
-               % [ % [ % [ (* Hc0, Hsplit *) @daemon | (*Hrs3*) @daemon ] | % ]
-                     | % ] | (* Hc0 *) @daemon ] ] | >Hta % ]
-             | * #tb * whd in ⊢ (%→?); #Hta'
-               whd in ⊢ (%→?); #Htb
-               cases (Hta' c0 ?)
-               [ #Hfalse @False_ind >Hfalse in Hc0;
-                 #Hc0 destruct (Hc0)
-               | >Hta % ]
-]]]]
-qed.
\ No newline at end of file
+]]]]]
+qed.       
+           
+axiom sem_compare : Realize ? compare R_compare.
\ No newline at end of file