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.
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 //
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.
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.
*)
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.
λ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.
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.
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.
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 ?))
qed.
definition compare ≝
- whileTM ? comp_step (inr … (inl … (inr … 0))).
+ whileTM ? comp_step (inr … (inl … (inr … start_nop))).
(*
definition R_compare :=
definition tc_states ≝ initN 3.
-definition tc_true ≝ 1.
+definition tc_start : tc_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 3 (refl …)).
+definition tc_true : tc_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 3 (refl …)).
+definition tc_false : tc_states ≝ mk_Sig ?? 2 (leb_true_to_le 3 3 (refl …)).
definition test_char ≝
λalpha:FinSet.λtest:alpha→bool.
mk_TM alpha tc_states
(λp.let 〈q,a〉 ≝ p in
match a with
- [ None ⇒ 〈1, None ?〉
+ [ None ⇒ 〈tc_true, None ?〉
| Some a' ⇒
match test a' with
- [ true ⇒ 〈1,None ?〉
- | false ⇒ 〈2,None ?〉 ]])
- O (λx.notb (x == 0)).
+ [ true ⇒ 〈tc_true,None ?〉
+ | false ⇒ 〈tc_false,None ?〉 ]])
+ tc_start (λx.notb (x == tc_start)).
definition Rtc_true ≝
λalpha,test,t1,t2.
lemma tc_q0_q1 :
∀alpha,test,ls,a0,rs. test a0 = true →
step alpha (test_char alpha test)
- (mk_config ?? 0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (test_char alpha test)) 1
+ (mk_config ?? tc_start (midtape … ls a0 rs)) =
+ mk_config alpha (states ? (test_char alpha test)) tc_true
(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 tc_q0_q2 :
∀alpha,test,ls,a0,rs. test a0 = false →
step alpha (test_char alpha test)
- (mk_config ?? 0 (midtape … ls a0 rs)) =
- mk_config alpha (states ? (test_char alpha test)) 2
+ (mk_config ?? tc_start (midtape … ls a0 rs)) =
+ mk_config alpha (states ? (test_char alpha test)) tc_false
(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 sem_test_char :
∀alpha,test.
accRealize alpha (test_char alpha test)
- 1 (Rtc_true alpha test) (Rtc_false alpha test).
+ tc_true (Rtc_true alpha test) (Rtc_false alpha test).
#alpha #test *
[ @(ex_intro ?? 2)
- @(ex_intro ?? (mk_config ?? 1 (niltape ?))) %
+ @(ex_intro ?? (mk_config ?? tc_true (niltape ?))) %
[ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (leftof ? a al)))
+| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_true (leftof ? a al)))
% [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ]
-| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? 1 (rightof ? a al)))
+| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_true (rightof ? a al)))
% [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ]
| #ls #c #rs @(ex_intro ?? 2)
cases (true_or_false (test c)) #Htest
- [ @(ex_intro ?? (mk_config ?? 1 ?))
+ [ @(ex_intro ?? (mk_config ?? tc_true ?))
[| %
[ %
[ whd in ⊢ (??%?); >tc_q0_q1 //
| #_ #c0 #Hc0 % // normalize in Hc0; destruct // ]
| * #Hfalse @False_ind @Hfalse % ]
]
- | @(ex_intro ?? (mk_config ?? 2 (midtape ? ls c rs)))
+ | @(ex_intro ?? (mk_config ?? tc_false (midtape ? ls c rs)))
%
[ %
[ whd in ⊢ (??%?); >tc_q0_q2 //
- | #Hfalse destruct (Hfalse) ]
+ | whd in ⊢ ((??%%)→?); #Hfalse destruct (Hfalse) ]
| #_ #c0 #Hc0 % // normalize in Hc0; destruct (Hc0) //
]
]