X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Fmarks.ma;h=829f56dc57ea7c1a2f80eef71df3a2d60a399ce4;hb=31cb2f0b374657eb5acb95708443e2c1b8481891;hp=4b66e6af524fe89362fdc19d455a3ac80b37aa53;hpb=35e0d9c8c27601b93413f4e9cbf1558404f24a41;p=helm.git diff --git a/matita/matita/lib/turing/universal/marks.ma b/matita/matita/lib/turing/universal/marks.ma index 4b66e6af5..829f56dc5 100644 --- a/matita/matita/lib/turing/universal/marks.ma +++ b/matita/matita/lib/turing/universal/marks.ma @@ -14,10 +14,9 @@ *) -include "turing/while_machine.ma". include "turing/if_machine.ma". +include "turing/basic_machines.ma". include "turing/universal/alphabet.ma". -include "turing/universal/tests.ma". (* ADVANCE TO MARK (right) @@ -30,17 +29,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 +60,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 +121,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 +193,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. @@ -216,77 +224,6 @@ lemma sem_mark : @ex_intro [| % [ % | #ls0 #c0 #b0 #rs0 #H1 destruct (H1) % ] ] ] qed. -(* MOVE RIGHT - - moves the head one step to the right - -*) - -definition move_states ≝ initN 2. - -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). - -definition R_move_r ≝ λalpha,t1,t2. - ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs). - -lemma sem_move_r : - ∀alpha.Realize ? (move_r alpha) (R_move_r alpha). -#alpha #intape @(ex_intro ?? 2) cases intape -[ @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -| #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases rs0 // ] ] ] -qed. - -(* MOVE LEFT - - moves the head one step to the right - -*) - -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). - -definition R_move_l ≝ λalpha,t1,t2. - ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (tail ? ls) (option_hd ? ls) (c::rs). - -lemma sem_move_l : - ∀alpha.Realize ? (move_l alpha) (R_move_l alpha). -#alpha #intape @(ex_intro ?? 2) cases intape -[ @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -|#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] -| #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases ls0 // ] ] ] -qed. (* MOVE RIGHT AND MARK machine @@ -297,18 +234,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 +278,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. @@ -498,10 +443,6 @@ qed. ^ *) -definition is_marked ≝ - λalpha.λp:FinProd … alpha FinBool. - let 〈x,b〉 ≝ p in b. - definition adv_both_marks ≝ λalpha.seq ? (adv_mark_r alpha) (seq ? (move_l ?) @@ -650,7 +591,7 @@ cases (sem_if ? (test_char ? (λx.x == c)) … tc_true | * #Hx0 #Houtc %2 % [ <(\P Hx) in Hx0; #Hx0 lapply (\Pf Hx0) @not_to_not #Hx' >Hx' % | >Houtc % ] - | (* members of lists are invariant under reverse *) @daemon ] + | #x #membx @Hl1 <(reverse_reverse …l1) @memb_reverse @membx ] | %2 % [ @(\Pf Hc) ] >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta #Hx #Hta >Hx in Hc;#Hc destruct (Hc) ] @@ -678,7 +619,8 @@ 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. @@ -686,7 +628,7 @@ 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) → @@ -695,7 +637,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. @@ -703,13 +645,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) % @@ -737,13 +680,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) + ] ] ] ] @@ -756,7 +710,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 := @@ -807,16 +761,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) @@ -832,32 +786,6 @@ definition R_compare := reverse ? la@ls) 〈d',false〉 (lc@〈comma,false〉::l2)). -lemma list_ind2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop. - length ? l1 = length ? l2 → - (P [] []) → - (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → - P l1 l2. -#T1 #T2 #l1 #l2 #P #Hl #Pnil #Pcons -generalize in match Hl; generalize in match l2; -elim l1 -[#l2 cases l2 // normalize #t2 #tl2 #H destruct -|#t1 #tl1 #IH #l2 cases l2 - [normalize #H destruct - |#t2 #tl2 #H @Pcons @IH normalize in H; destruct // ] -] -qed. - -lemma list_cases_2 : - ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop. - length ? l1 = length ? l2 → - (l1 = [] → l2 = [] → P) → - (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P. -#T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl) -[ #Pnil #Pcons @Pnil // -| #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ] -qed. - lemma wsem_compare : WRealize ? compare R_compare. #t #i #outc #Hloop lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] @@ -873,7 +801,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 #_ % @@ -893,7 +821,7 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] |#b #b0 #bs #b0s #l1 #l2 #Hlen #Hbs1 #Hb0s1 #Hbs2 #Hb0s2 #Hl1 #Heq destruct (Heq) #_ #Hrs cases Hleft -Hleft [2: * >Hc' #Hfalse @False_ind destruct ] * #_ - @(list_cases_2 … Hlen) + @(list_cases2 … 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 #_ #_ @@ -912,7 +840,7 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] | @Hl1 ] | * #b' #bitb' * #b0' #bitb0' #bs' #b0s' #Hbs #Hb0s generalize in match Hrs; >Hbs in ⊢ (%→?); >Hb0s in ⊢ (%→?); - cut (is_bit b' = true ∧ is_bit b0' = true ∧ + 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 ] @@ -970,5 +898,5 @@ lapply (sem_while ?????? sem_comp_step t i outc Hloop) [%] ] ]]]]] qed. - + axiom sem_compare : Realize ? compare R_compare. \ No newline at end of file