X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fbasic_machines.ma;h=c1559702a59770c27dc68bbfcf98e2d50f6aaa46;hb=48c011f52853dd106dbf9cbbd1b9da61277fba3b;hp=5612353a81f231e1f34898c557da87bee2559bd8;hpb=48ea5e91651eef80927defbdd91af0e9e3892999;p=helm.git diff --git a/matita/matita/lib/turing/basic_machines.ma b/matita/matita/lib/turing/basic_machines.ma index 5612353a8..c1559702a 100644 --- a/matita/matita/lib/turing/basic_machines.ma +++ b/matita/matita/lib/turing/basic_machines.ma @@ -21,8 +21,8 @@ definition write ≝ λalpha,c. mk_TM alpha write_states (λp.let 〈q,a〉 ≝ p in match pi1 … q with - [ O ⇒ 〈wr1,Some ? 〈c,N〉〉 - | S _ ⇒ 〈wr1,None ?〉 ]) + [ O ⇒ 〈wr1,Some ? c,N〉 + | S _ ⇒ 〈wr1,None ?,N〉 ]) wr0 (λx.x == wr1). definition R_write ≝ λalpha,c,t1,t2. @@ -32,7 +32,35 @@ lemma sem_write : ∀alpha,c.Realize ? (write alpha c) (R_write alpha c). #alpha #c #t @(ex_intro … 2) @ex_intro [|% [% |#ls #c #rs #Ht >Ht % ] ] qed. - + +definition R_write_strong ≝ λalpha,c,t1,t2. + t2 = midtape alpha (left ? t1) c (right ? t1). + +lemma sem_write_strong : ∀alpha,c.Realize ? (write alpha c) (R_write_strong alpha c). +#alpha #c #t @(ex_intro … 2) @ex_intro + [|% [% |cases t normalize // ] ] +qed. + +(***************************** replace a with f a *****************************) + +definition writef ≝ λalpha,f. + mk_TM alpha write_states + (λp.let 〈q,a〉 ≝ p in + match pi1 … q with + [ O ⇒ 〈wr1,Some ? (f a),N〉 + | S _ ⇒ 〈wr1,None ?,N〉 ]) + wr0 (λx.x == wr1). + +definition R_writef ≝ λalpha,f,t1,t2. + ∀c. current ? t1 = c → + t2 = midtape alpha (left ? t1) (f c) (right ? t1). + +lemma sem_writef : ∀alpha,f. + writef alpha f ⊨ R_writef alpha f. +#alpha #f #t @(ex_intro … 2) @ex_intro + [|% [% |cases t normalize // ] ] +qed. + (******************** moves the head one step to the right ********************) definition move_states ≝ initN 2. @@ -43,29 +71,35 @@ definition move_r ≝ λalpha:FinSet.mk_TM alpha move_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈move1,None ?〉 + [ None ⇒ 〈move1,None ?,N〉 | Some a' ⇒ match (pi1 … q) with - [ O ⇒ 〈move1,Some ? 〈a',R〉〉 - | S q ⇒ 〈move1,None ?〉 ] ]) + [ O ⇒ 〈move1,Some ? a',R〉 + | S q ⇒ 〈move1,None ?,N〉 ] ]) move0 (λq.q == move1). definition R_move_r ≝ λalpha,t1,t2. + (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs). + t1 = midtape alpha ls c rs → + t2 = mk_tape ? (c::ls) (option_hd ? rs) (tail ? rs). +(* + (current ? t1 = None ? ∧ 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 ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] |#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] |#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] | #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases rs0 // ] ] ] + @ex_intro [| % [ % | % [whd in ⊢ ((??%?)→?); #H destruct] + #ls1 #c1 #rs1 #H destruct cases rs1 // ] ] ] qed. (******************** moves the head one step to the left *********************) @@ -74,29 +108,54 @@ definition move_l ≝ λalpha:FinSet.mk_TM alpha move_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈move1,None ?〉 + [ None ⇒ 〈move1,None ?,N〉 | Some a' ⇒ match pi1 … q with - [ O ⇒ 〈move1,Some ? 〈a',L〉〉 - | S q ⇒ 〈move1,None ?〉 ] ]) + [ O ⇒ 〈move1,Some ? a',L〉 + | S q ⇒ 〈move1,None ?,N〉 ] ]) move0 (λq.q == move1). - + definition R_move_l ≝ λalpha,t1,t2. + (current ? t1 = None ? → t1 = t2) ∧ ∀ls,c,rs. - t1 = midtape alpha ls c rs → - t2 = mk_tape ? (tail ? ls) (option_hd ? 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 ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] |#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] |#a #al @ex_intro - [| % [ % | #ls #c #rs #Hfalse destruct ] ] + [| % [ % | % // #ls #c #rs #H destruct ] ] | #ls #c #rs - @ex_intro [| % [ % | #ls0 #c0 #rs0 #H1 destruct (H1) - cases ls0 // ] ] ] + @ex_intro [| % [ % | % [whd in ⊢ ((??%?)→?); #H destruct] + #ls1 #c1 #rs1 #H destruct cases ls1 // ] ] ] +qed. + +(* a slightly different move machine. *) +definition smove_states ≝ initN 2. + +definition smove0 : smove_states ≝ mk_Sig ?? 0 (leb_true_to_le 1 2 (refl …)). +definition smove1 : smove_states ≝ mk_Sig ?? 1 (leb_true_to_le 2 2 (refl …)). + +definition trans_smove ≝ + λsig,D. + λp:smove_states × (option sig). + let 〈q,a〉 ≝ p in match (pi1 … q) with + [ O ⇒ 〈smove1,None sig, D〉 + | S _ ⇒ 〈smove1,None sig, N〉 ]. + +definition move ≝ + λsig,D.mk_TM sig smove_states (trans_smove sig D) smove0 (λq.q == smove1). + +definition Rmove ≝ + λalpha,D,t1,t2. t2 = tape_move alpha t1 D. + +lemma sem_move_single : + ∀alpha,D.move alpha D ⊨ Rmove alpha D. +#alpha #D #int %{2} %{(mk_config ? smove_states smove1 ?)} [| % % ] qed. (********************************* test char **********************************) @@ -118,20 +177,20 @@ definition test_char ≝ mk_TM alpha tc_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈tc_true, None ?〉 + [ None ⇒ 〈tc_false, None ?,N〉 | Some a' ⇒ match test a' with - [ true ⇒ 〈tc_true,None ?〉 - | false ⇒ 〈tc_false,None ?〉 ]]) + [ true ⇒ 〈tc_true,None ?,N〉 + | false ⇒ 〈tc_false,None ?,N〉 ]]) tc_start (λx.notb (x == tc_start)). definition Rtc_true ≝ λalpha,test,t1,t2. - ∀c. current alpha t1 = Some ? c → test c = true ∧ t2 = t1. + (∃c. current alpha t1 = Some ? c ∧ test c = true) ∧ t2 = t1. definition Rtc_false ≝ λalpha,test,t1,t2. - ∀c. current alpha t1 = Some ? c → test c = false ∧ t2 = t1. + (∀c. current alpha t1 = Some ? c → test c = false) ∧ t2 = t1. lemma tc_q0_q1 : ∀alpha,test,ls,a0,rs. test a0 = true → @@ -159,19 +218,19 @@ lemma sem_test_char : tc_true (Rtc_true alpha test) (Rtc_false alpha test). #alpha #test * [ @(ex_intro ?? 2) - @(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 ?? tc_true (leftof ? a al))) - % [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ] -| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_true (rightof ? a al))) - % [ % // #_ #c normalize #Hfalse destruct | #_ #c normalize #Hfalse destruct (Hfalse) ] + @(ex_intro ?? (mk_config ?? tc_false (niltape ?))) % + [ % // normalize #Hfalse destruct | #_ normalize % // #c #Hfalse destruct (Hfalse) ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_false (leftof ? a al))) + % [ % // normalize #Hfalse destruct | #_ normalize % // #c #Hfalse destruct (Hfalse) ] +| #a #al @(ex_intro ?? 2) @(ex_intro ?? (mk_config ?? tc_false (rightof ? a al))) + % [ % // normalize #Hfalse destruct | #_ normalize % // #c #Hfalse destruct (Hfalse) ] | #ls #c #rs @(ex_intro ?? 2) cases (true_or_false (test c)) #Htest [ @(ex_intro ?? (mk_config ?? tc_true ?)) [| % [ % [ whd in ⊢ (??%?); >tc_q0_q1 // - | #_ #c0 #Hc0 % // normalize in Hc0; destruct // ] + | #_ % // @(ex_intro … c) /2/] | * #Hfalse @False_ind @Hfalse % ] ] | @(ex_intro ?? (mk_config ?? tc_false (midtape ? ls c rs))) @@ -179,9 +238,367 @@ lemma sem_test_char : [ % [ whd in ⊢ (??%?); >tc_q0_q2 // | whd in ⊢ ((??%%)→?); #Hfalse destruct (Hfalse) ] - | #_ #c0 #Hc0 % // normalize in Hc0; destruct (Hc0) // + | #_ % // #c0 whd in ⊢ ((??%?)→?); #Hc0 destruct (Hc0) // ] ] ] qed. +lemma test_char_inv : + ∀sig.∀P:tape sig → Prop.∀f,t,t0.P t → Rtc_true sig f t t0 → P t0. +#sig #P #f #t #t0 #HPt * #_ // +qed. + +definition test_null ≝ λalpha.test_char alpha (λ_.true). + +definition R_test_null_true ≝ λalpha,t1,t2. + current alpha t1 ≠ None ? ∧ t1 = t2. + +definition R_test_null_false ≝ λalpha,t1,t2. + current alpha t1 = None ? ∧ t1 = t2. + +lemma sem_test_null : ∀alpha. + test_null alpha ⊨ [ tc_true : R_test_null_true alpha, R_test_null_false alpha]. +#alpha #t1 cases (sem_test_char alpha (λ_.true) t1) #k * #outc * * #Hloop #Htrue +#Hfalse %{k} %{outc} % [ % +[ @Hloop +| #Houtc cases (Htrue ?) [| @Houtc] * #c * #Hcurt1 #_ #Houtc1 % + [ >Hcurt1 % #H destruct (H) |