X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Fbasic_machines.ma;h=4f1977dd47ec3ea8b7a2d6afaac77acde3f87e35;hb=b23a40026442aaad0fdba4590c720a33672ec4ce;hp=1c1ce3020b824d3c37c1bb48f7581c060deec56f;hpb=fd282412fff8f2529bb1dfb22a684f5c25af37cb;p=helm.git diff --git a/matita/matita/lib/turing/basic_machines.ma b/matita/matita/lib/turing/basic_machines.ma index 1c1ce3020..4f1977dd4 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. @@ -43,10 +43,10 @@ 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. @@ -80,10 +80,10 @@ 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. @@ -125,11 +125,11 @@ definition test_char ≝ mk_TM alpha tc_states (λp.let 〈q,a〉 ≝ p in match a with - [ None ⇒ 〈tc_false, 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 ≝ @@ -192,6 +192,32 @@ lemma sem_test_char : ] 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) |