X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Funi_step.ma;h=ea8e991e28b59ced2ed2aca9007f651edd6983cf;hb=9a7e0697f0b305e4ba26b67c37681c434709509a;hp=6a6bdc07cd297b395b80e9a796c71c7846f7e8c3;hpb=356324638c26395cb2ff465b615640c857db05be;p=helm.git diff --git a/matita/matita/lib/turing/universal/uni_step.ma b/matita/matita/lib/turing/universal/uni_step.ma index 6a6bdc07c..ea8e991e2 100644 --- a/matita/matita/lib/turing/universal/uni_step.ma +++ b/matita/matita/lib/turing/universal/uni_step.ma @@ -244,32 +244,19 @@ definition exec_move ≝ definition map_move ≝ λc,mv.match c with [ null ⇒ None ? | _ ⇒ Some ? 〈c,false,move_of_unialpha mv〉 ]. -definition current_of_alpha ≝ λc:STape. - match \fst c with [ null ⇒ None ? | _ ⇒ Some ? c ]. - -definition legal_tape ≝ λls,c,rs. - let t ≝ mk_tape STape ls (current_of_alpha c) rs in - left ? t = ls ∧ right ? t = rs ∧ current ? t = current_of_alpha c. - -lemma legal_tape_cases : - ∀ls,c,rs.legal_tape ls c rs → - \fst c ≠ null ∨ (\fst c = null ∧ ls = []) ∨ (\fst c = null ∧ rs = []). -#ls #c #rs cases c #c0 #bc0 cases c0 -[ #c1 normalize #_ % % % #Hfalse destruct (Hfalse) -| cases ls - [ #_ % %2 % % - | #l0 #ls0 cases rs - [ #_ %2 % % - | #r0 #rs0 normalize * * #Hls #Hrs destruct (Hrs) ] - ] -|*: #_ % % % #Hfalse destruct (Hfalse) ] -qed. - +(* - aggiungere a legal_tape le condizioni + only_bits ls, rs; bit_or_null c + - ci vuole un lemma che dimostri + bit_or_null c1 = true bit_or_null mv = true + mv ≠ null → c1 ≠ null + dal fatto che c1 e mv sono contenuti nella table + *) definition R_exec_move ≝ λt1,t2. ∀n,curconfig,ls,rs,c0,c1,s0,s1,table1,newconfig,mv,table2. table_TM n (table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2) → - no_marks curconfig → only_bits (curconfig@[〈s0,false〉]) → only_bits (〈s1,false〉::newconfig) → - no_nulls ls → no_nulls rs → no_marks ls → no_marks rs → + no_marks curconfig → only_bits (curconfig@[〈s0,false〉]) → + only_bits (〈s1,false〉::newconfig) → bit_or_null c1 = true → + |curconfig| = |newconfig| → legal_tape ls 〈c0,false〉 rs → t1 = midtape STape (〈c0,false〉::curconfig@〈s0,false〉::〈grid,false〉::ls) 〈grid,false〉 (table1@〈comma,true〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) → @@ -279,7 +266,7 @@ definition R_exec_move ≝ λt1,t2. (〈s1,false〉::newconfig@〈c2,false〉::〈grid,false〉:: table1@〈comma,false〉::〈s1,false〉::newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs1) ∧ lift_tape ls1 〈c2,false〉 rs1 = - tape_move STape t1' (map_move c1 mv). + tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1. (* move the following 2 lemmata to mono.ma *) lemma tape_move_left_eq : @@ -296,6 +283,48 @@ lemma tape_move_right_eq : // qed. +lemma lift_tape_not_null : + ∀ls,c,bc,rs.c ≠ null → lift_tape ls 〈c,bc〉 rs = midtape ? ls 〈c,bc〉 rs. +#ls #c #bc #rs cases c // +#Hfalse @False_ind /2/ +qed. + +lemma merge_char_not_null : + ∀c1,c2.c1 ≠ null → merge_char c1 c2 ≠ null. +#c1 #c2 @not_to_not cases c2 +[ #c1' normalize #Hfalse destruct (Hfalse) +| normalize // +| *: normalize #Hfalse destruct (Hfalse) +] +qed. + +lemma merge_char_null : ∀c.merge_char null c = c. +* // +qed. + +lemma merge_char_cases : ∀c1,c2.merge_char c1 c2 = c1 ∨ merge_char c1 c2 = c2. +#c1 * +[ #c1' %2 % +| % % +| *: %2 % ] +qed. + +(* lemma merge_char_c_bit : + ∀c1,c2.is_bit c2 = true → merge_char c1 c2 = c2. +#c1 * +[ #c2' #_ % +|*: normalize #Hfalse destruct (Hfalse) ] +qed. + +lemma merge_char_c_bit : + ∀c1,c2.is_null c2 = true → merge_char c1 c2 = c1. +#c1 * +[ #c2' #_ % +|*: normalize #Hfalse destruct (Hfalse) ] +qed. + +*) + lemma sem_exec_move : Realize ? exec_move R_exec_move. #intape cases (sem_seq … sem_init_copy @@ -304,24 +333,36 @@ cases (sem_seq … sem_init_copy #k * #outc * #Hloop #HR @(ex_intro ?? k) @(ex_intro ?? outc) % [ @Hloop ] -Hloop #n #curconfig #ls #rs #c0 #c1 #s0 #s1 #table1 #newconfig #mv #table2 -#Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hls #Hrs #Hls1 #Hrs1 #Htape #Hintape #t1' #Ht1' +#Htable #Hcurconfig1 #Hcurconfig2 #Hnewconfig #Hc1 #Hlen #Htape #Hintape #t1' #Ht1' cases HR -HR #ta * whd in ⊢ (%→?); #Hta lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1 (newconfig@〈c1,false〉::〈comma,false〉::〈mv,false〉::table2@〈grid,false〉::rs) … Hintape) -Hta [ (*Hcurconfig2*) @daemon | (*Htable*) @daemon -| (*FIXME*) @daemon -| (*FIXME*) @daemon +| (*bit_or_null c0 = true *) @daemon +| (*Hcurconfig1*) @daemon | #Hta * #tb * whd in ⊢ (%→?); #Htb lapply (Htb (〈grid,false〉::ls) s0 s1 c0 c1 (〈mv,false〉::table2@〈grid,false〉::rs) newconfig (〈comma,false〉::reverse ? table1) curconfig Hta ????????) -Htb - [9:|*:(* rivedere le precondizioni *) @daemon ] + [9:|*:(* bit_or_null c0,c1; |curconfig| = |newconfig|*) @daemon ] #Htb * #tc * whd in ⊢ (%→?); #Htc lapply (Htc … Htb) -Htc whd in ⊢(???(??%%%)→?);#Htc whd in ⊢ (%→?); #Houtc whd in Htc:(???%); whd in Htc:(???(??%%%)); lapply (Houtc rs n (〈comma,false〉::〈c1,false〉::reverse ? newconfig@〈s1,false〉::〈comma,false〉::reverse ? table1) - mv table2 (merge_char c0 c1) (reverse ? newconfig@[〈s1,false〉]) ls ????? - Hls Hrs Hls1 Hrs1 ??) - [3: >Htc @(eq_f3 … (midtape ?)) + mv table2 (merge_char c0 c1) (reverse ? newconfig@[〈s1,false〉]) ls ????????) + [3: cases Htape -Htape * * #Hnomarks #Hbits #Hc0 #Hlsrs % [ % [ % + [ #x #Hx cases (orb_true_l … Hx) #Hx' + [ >(\P Hx') % + | @Hnomarks @memb_cons // ] + | @Hbits ] + | cases (merge_char_cases c0 c1) #Hmerge >Hmerge // ] + | cases (true_or_false (c0 == null)) #Hc0' + [ cases Hlsrs -Hlsrs + [ * + [ >(\P Hc0') * #Hfalse @False_ind /2/ + | #Hlsnil % %2 // ] + | #Hrsnil %2 // ] + | % % @merge_char_not_null @(\Pf Hc0') ] ] + |4:>Htc @(eq_f3 … (midtape ?)) [ @eq_f @eq_f >associative_append >associative_append % | % | % ] @@ -331,14 +372,51 @@ lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1 >associative_append >associative_append >associative_append >associative_append >associative_append @Htable - | (* only bits or nulls c1,c2 → only bits or nulls (merge c1 c2) *) @daemon - | (* add to hyps? *) @daemon - | (* bit_or_null c1,c2 → bit_or_null (merge_char c1 c2) *) @daemon - | -Houtc * #ls1 * #rs1 * #newc * #Houtc * + | (* well formedness of table *) @daemon + | (* Hnewconfig *) @daemon + | (* bit_or_null mv = true (well formedness of table) *) @daemon + | -Houtc * #ls1 * #rs1 * #newc * #Hnewtapelegal * #Houtc * [ * [ * #Hmv #Htapemove @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) % + [ % + [ >Houtc -Houtc >reverse_append + >reverse_reverse >reverse_single @eq_f + >reverse_cons >reverse_cons >reverse_append >reverse_cons + >reverse_cons >reverse_reverse >reverse_reverse + >associative_append >associative_append + >associative_append >associative_append + >associative_append >associative_append % + | >Hmv >Ht1' >Htapemove + (* mv = bit false -→ c1 = bit ? *) + cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1 + >Hc1 >tape_move_left_eq >(legal_tape_left … Htape) + >(legal_tape_right … Htape) % + ] + | // + ] + | * #Hmv #Htapemove + @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) % + [ % + [ >Houtc -Houtc >reverse_append + >reverse_reverse >reverse_single @eq_f + >reverse_cons >reverse_cons >reverse_append >reverse_cons + >reverse_cons >reverse_reverse >reverse_reverse + >associative_append >associative_append + >associative_append >associative_append + >associative_append >associative_append % + |>Hmv >Ht1' >Htapemove + cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1 + >Hc1 >tape_move_right_eq >(legal_tape_left … Htape) + >(legal_tape_right … Htape) % + ] + | // + ] + ] + | * * * #Hmv #Hlseq #Hrseq #Hnewc + @(ex_intro ?? ls1) @(ex_intro ?? rs1) @(ex_intro ?? newc) % + [ % [ >Houtc -Houtc >reverse_append >reverse_reverse >reverse_single @eq_f >reverse_cons >reverse_cons >reverse_append >reverse_cons @@ -346,69 +424,156 @@ lapply (Hta (〈c0,false〉::curconfig) table1 s0 ls s1 >associative_append >associative_append >associative_append >associative_append >associative_append >associative_append % - | >Hmv >Ht1' >Htapemove - (* mv = bit false -→ c1 = bit ? - whd in Htapemove:(???%); whd in ⊢ (???%); - whd in match (lift_tape ???) in ⊢ (???%); - >Htapemove *) - cut (∃c1'.c1 = bit c1') [ @daemon ] * #c1' #Hc1 - >Hc1 >tape_move_left_eq cases c0 - [ #c0' % - | cases ls - [ cases rs - [ % - | #r0 #rs0 % ] - | #l0 #ls0 cases rs - [ % - | #r0 #rs0 - - ls #... null # ... # rs - ls #... c # ... # rs - - ∃t.left ? t = ls, right ? t = rs, current t = [[ c ]] - - % ] - - - @eq_f3 - [ - - whd in match (merge_char ??); - whd in match (map_move ??); - + |>Hmv >Ht1' cases c1 in Hnewc; + [ #c1' whd in ⊢ (??%?→?);#Hnewc Hlseq >Hrseq whd in ⊢ (??%%); + >(legal_tape_left … Htape) >(legal_tape_right … Htape) % + | whd in ⊢ (??%?→?); #Hnewc >Hnewc >Hlseq >Hrseq % + |*: whd in ⊢ (??%?→?);#Hnewc Hlseq >Hrseq whd in ⊢ (??%%); + >(legal_tape_left … Htape) >(legal_tape_right … Htape) % + ] ] - | - + | // + ] + ] + ] +] +qed. + +(* +if is_false(current) (* current state is not final *) + then init_match; + match_tuple; + if is_marked(current) = false (* match ok *) + then + exec_move + move_r; + else sink; + else nop; +*) + +definition uni_step ≝ + ifTM ? (test_char STape (λc.\fst c == bit false)) + (single_finalTM ? (seq ? init_match + (seq ? match_tuple + (ifTM ? (test_char ? (λc.¬is_marked ? c)) + (seq ? exec_move (move_r …)) + (nop ?) (* sink *) + tc_true)))) + (nop ?) + tc_true. +definition R_uni_step_true ≝ λt1,t2. + ∀n,t0,table,s0,s1,c0,c1,ls,rs,curconfig,newconfig,mv. + table_TM (S n) (〈t0,false〉::table) → + match_in_table (S n) (〈s0,false〉::curconfig) 〈c0,false〉 + (〈s1,false〉::newconfig) 〈c1,false〉 〈mv,false〉 (〈t0,false〉::table) → + legal_tape ls 〈c0,false〉 rs → + t1 = midtape STape (〈grid,false〉::ls) 〈s0,false〉 + (curconfig@〈c0,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs) → + ∀t1'.t1' = lift_tape ls 〈c0,false〉 rs → + s0 = bit false ∧ + ∃ls1,rs1,c2. + (t2 = midtape STape (〈grid,false〉::ls1) 〈s1,false〉 + (newconfig@〈c2,false〉::〈grid,false〉::〈t0,false〉::table@〈grid,false〉::rs1) ∧ + lift_tape ls1 〈c2,false〉 rs1 = + tape_move STape t1' (map_move c1 mv) ∧ legal_tape ls1 〈c2,false〉 rs1). + +definition R_uni_step_false ≝ λt1,t2. + ∀b. current STape t1 = Some ? 〈bit b,false〉 → b = true ∧ t2 = t1. - >Heqcurconfig ->append_cons in Hintape; >associative_append -cases (Hta - - cases (Hta … Hintape) -Hta -[ * normalize in ⊢ (%→?); #Hfalse destruct (Hfalse) ] -* #_ #Hta lapply (Hta ? 〈comma,true〉 … (refl ??) (refl ??) ?) -[ @daemon ] -Hta #Hta -* #tb * whd in ⊢ (%→?); #Htb lapply (Htb … Hta) - -definition move_of_unialpha ≝ - λc.match c with - [ bit x ⇒ match x with [ true ⇒ R | false ⇒ L ] - | _ ⇒ N ]. - -definition R_uni_step ≝ λt1,t2. - ∀n,table,c,c1,ls,rs,curs,curc,news,newc,mv. - table_TM n table → - match_in_table (〈c,false〉::curs@[〈curc,false〉]) - (〈c1,false〉::news@[〈newc,false〉]) mv table → - t1 = midtape STape (〈grid,false〉::ls) 〈c,false〉 - (curs@〈curc,false〉::〈grid,false〉::table@〈grid,false〉::rs) → - ∀t1',ls1,rs1.t1' = lift_tape ls 〈curc,false〉 rs → - (t2 = midtape STape (〈grid,false〉::ls1) 〈c1,false〉 - (news@〈newc,false〉::〈grid,false〉::table@〈grid,false〉::rs1) ∧ - lift_tape ls1 〈newc,false〉 rs1 = - tape_move STape t1' (Some ? 〈〈newc,false〉,move_of_unialpha mv〉)). - -definition no_nulls ≝ - λl:list STape.∀x.memb ? x l = true → is_null (\fst x) = false. - +axiom sem_match_tuple : Realize ? match_tuple R_match_tuple. + +lemma sem_uni_step : + accRealize ? uni_step (inr … (inl … (inr … 0))) + R_uni_step_true R_uni_step_false. +@(acc_sem_if_app … (sem_test_char ? (λc:STape.\fst c == bit false)) + (sem_seq … sem_init_match + (sem_seq … sem_match_tuple + (sem_if … (sem_test_char ? (λc.¬is_marked ? c)) + (sem_seq … sem_exec_move (sem_move_r …)) + (sem_nop …)))) + (sem_nop …) + …) +[ #intape #outtape + #ta whd in ⊢ (%→?); #Hta #HR + #n #t0 #table #s0 #s1 #c0 #c1 #ls #rs #curconfig #newconfig #mv + #Htable #Hmatch #Htape #Hintape #t1' #Ht1' + >Hintape in Hta; #Hta cases (Hta ? (refl ??)) -Hta + #Hs0 lapply (\P Hs0) -Hs0 #Hs0 #Hta % // + cases HR -HR + #tb * whd in ⊢ (%→?); #Htb + lapply (Htb (〈grid,false〉::ls) (curconfig@[〈c0,false〉]) (table@〈grid,false〉::rs) s0 t0 ???) + [ >Hta >associative_append % + | (* utilizzare Hmatch + cases (match_in_table_to_tuple … Hmatch Htable) + ma serve l'iniettività di mk_tuple... + *) @daemon + | (* idem *) @daemon + | -Hta -Htb #Htb * + #tc * whd in ⊢ (%→?); #Htc cases (Htc … Htable … Htb) -Htb -Htc + [| * #Hcurrent #Hfalse @False_ind + (* absurd by Hmatch *) @daemon + | >Hs0 % + | (* da decidere se aggiungere un'assunzione o utilizzare Hmatch *) @daemon + | (* Htable *) @daemon + | (* Htable, Hmatch → |config| = n + necessaria modifica in R_match_tuple, le dimensioni non corrispondono + *) @daemon ] + * #table1 * #newc * #mv1 * #table2 * #Htableeq #Htc * + [ * #td * whd in ⊢ (%→?); >Htc -Htc #Htd + cases (Htd ? (refl ??)) #_ -Htd + cut (newc = 〈s1,false〉::newconfig@[〈c1,false〉]) [@daemon] #Hnewc + >Hnewc #Htd + cut (mv1 = 〈\fst mv1,false〉) + [ >(eq_pair_fst_snd … mv1) @eq_f (*Htable, Htableeq*) @daemon ] #Hmv1 + * #te * whd in ⊢ (%→?); #Hte + cut (td = midtape STape (〈c0,false〉::reverse STape curconfig@〈s0,false〉::〈grid,false〉::ls) + 〈grid,false〉 + ((table1@〈s0,false〉::curconfig@[〈c0,false〉])@〈comma,true〉::〈s1,false〉:: + newconfig@〈c1,false〉::〈comma,false〉::〈\fst mv1,false〉::table2@〈grid,false〉::rs)) + [ >Htd @eq_f3 // + [ >reverse_append >reverse_single % + | >associative_append >associative_append normalize + >associative_append >associative_append Hnewc in Htableeq; + >associative_append >associative_append normalize + >associative_append >associative_append + #Htableeq Houttape @eq_f @eq_f @eq_f @eq_f + change with ((〈t0,false〉::table)@?) in ⊢ (???%); + >Htableeq >associative_append >associative_append + >associative_append normalize >associative_append + >associative_append normalize >Hnewc associative_append normalize >associative_append % + | >(?: mv = \fst mv1) [| (*Hmatch, Htableeq*) @daemon ] + @Hliftte + ] + | // + ] + ] + ] + | * #td * whd in ⊢ (%→%→?); >Htc #Htd + cases (Htd ? (refl ??)) normalize in ⊢ (%→?); + #Hfalse destruct (Hfalse) + ] + ] +| #t1 #t2 #t3 whd in ⊢ (%→%→?); #Ht1 #Ht2 + #b #Hb cases (Ht1 ? Hb) #Hb' #Ht3 >Ht2 % // + cases b in Hb'; normalize #H1 // +] +qed.