X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Fmarks.ma;h=70000e3996690e2837ead12ca3c932cfa581ae12;hb=5fc2b08d86038360e588b8fff333a623964efabe;hp=17992f95fca7cdef66e89e4f82c85abc174f338b;hpb=67948c54e4e20ef530d2d2115c9a056275360012;p=helm.git diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index 17992f95f..70000e399 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -30,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. @@ -57,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 // @@ -116,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. @@ -188,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. @@ -223,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. @@ -263,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. @@ -297,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. @@ -337,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. @@ -704,7 +723,7 @@ 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 ?)) @@ -769,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 :=