2 ||M|| This file is part of HELM, an Hypertextual, Electronic
3 ||A|| Library of Mathematics, developed at the Computer Science
4 ||T|| Department of the University of Bologna, Italy.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_____________________________________________________________*)
13 include "turing/multi_universal/unistep_aux.ma".
15 definition exec_move ≝
16 cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg.
18 definition R_exec_move ≝ λt1,t2:Vector (tape FSUnialpha) 3.
20 nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls1@[bar]) (None ?) [ ] →
21 nth prg ? t1 (niltape ?) = midtape FSUnialpha (ls2@[bar]) m rs2 →
22 only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) →
25 tape_move_mono ? (nth obj ? t1 (niltape ?))
26 〈char_to_bit_option c, char_to_move m〉 in
27 let next_c ≝ low_char' (current ? new_obj) in
28 let new_cfg ≝ midtape ? [ ] bar ((reverse ? ls1)@[next_c]) in
29 let new_prg ≝ midtape FSUnialpha [ ] bar ((reverse ? ls2)@m::rs2) in
30 t2 = Vector_of_list ? [new_obj;new_cfg;new_prg].
33 lemma sem_exec_move: exec_move ⊨ R_exec_move.
34 @(sem_seq_app ??????? sem_cfg_to_obj1
35 (sem_seq ?????? sem_tape_move_obj
36 (sem_seq ?????? (sem_restart_tape ???) sem_obj_to_cfg1))) //
37 #ta #tout * #t1 * #semM1 * #t2 * #semM2 * #t3 * #semM3 #semM4
38 #c #m #ls1 #ls2 #rs2 #Hcfg #Hprg #Honlybits #Hc #Hm
40 lapply (semM1 … Hcfg Hc) #Ht1
42 whd in semM2; >Ht1 in semM2; -Ht1
43 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
44 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
45 >Hprg #Ht2 lapply (Ht2 … (refl ??)) -Ht2
46 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
47 >nth_change_vec // >change_vec_commute [2:@eqb_false_to_not_eq %]
48 >change_vec_change_vec #Ht2
49 (* M3 = restart prg *)
50 whd in semM3; >Ht2 in semM3; #semM3 lapply (semM3 … (refl ??)); -semM3
51 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
52 >nth_change_vec_neq [2:@eqb_false_to_not_eq %] #Ht3
54 whd in semM4; >Ht3 in semM4;
55 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
56 >nth_change_vec [2:@leb_true_to_le %] #semM4 lapply (semM4 … (refl ??)) -semM4
57 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
58 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
59 >nth_change_vec [2:@leb_true_to_le %] #semM4 >(semM4 ?)
61 @(eq_vec … (niltape ?)) #i #lei2
62 cases (le_to_or_lt_eq … (le_S_S_to_le …lei2))
63 [#lei1 cases (le_to_or_lt_eq … (le_S_S_to_le …lei1))
64 [#lei0 lapply(le_n_O_to_eq … (le_S_S_to_le …lei0)) #eqi <eqi
65 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
66 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
67 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
68 >nth_change_vec [2:@leb_true_to_le %] %
69 |#Hi >Hi (* obj tape *)
70 >nth_change_vec [2:@leb_true_to_le %] whd in ⊢ (???%);
71 >reverse_cons >reverse_append >reverse_single
72 whd in match (option_hd ??); whd in match (tail ??);
75 |#Hi >Hi (* prg tape *)
76 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
77 >nth_change_vec [2:@leb_true_to_le %] whd in ⊢ (???%);
78 >Hprg whd in match (list_of_tape ??); >reverse_append
81 |#b #Hcurrent @(list_of_tape_write ? is_bit … (char_to_bit_option c) ? Honlybits)
83 [#b1 whd in ⊢ (??%?→?); #Hbit destruct (Hbit) %
84 |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
85 |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
87 |>(list_of_tape_move … (char_to_move m)) @current_in_list @Hcurrent
93 match_m cfg prg FSUnialpha 2 ·
94 restart_tape cfg 2 · mmove cfg ? 2 R · copy prg cfg FSUnialpha 2 ·
98 definition legal_tape ≝ λn,l,h,t.
100 nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) →
101 is_config n (bar::state@[char]) →
102 nth prg ? t1 (niltape ?) = midtape ? [ ] bar table →
103 bar::table = table_TM n l h → *)
105 definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3.
108 nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) →
109 is_config n (bar::state@[char]) →
111 nth prg ? t1 (niltape ?) = midtape ? [ ] bar table →
112 bar::table = table_TM n l h →
114 only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) →
115 let conf ≝ (bar::state@[char]) in
116 (∃ll,lr.bar::table = ll@conf@lr) →
118 ∃nstate,nchar,m,t. tuple_encoding n h t = (conf@nstate@[nchar;m]) ∧
121 tuple_encoding n h t = (conf@nstate@[nchar;m])→
124 tape_move_mono ? (nth obj ? t1 (niltape ?))
125 〈char_to_bit_option nchar,char_to_move m〉 in
126 let next_char ≝ low_char' (current ? new_obj) in
129 (change_vec ?? t1 (midtape ? [ ] bar (nstate@[next_char])) cfg)
132 lemma lt_obj : obj < 3. // qed.
133 lemma lt_cfg : cfg < 3. // qed.
134 lemma lt_prg : prg < 3. // qed.
136 definition R_copy_strict ≝
137 λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
138 ((current ? (nth src ? int (niltape ?)) = None ? ∨
139 current ? (nth dst ? int (niltape ?)) = None ?) → outt = int) ∧
140 (∀ls,x,x0,rs,ls0,rs0.
141 nth src ? int (niltape ?) = midtape sig ls x rs →
142 nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
144 (∃rs1,rs2.rs = rs1@rs2 ∧ |rs1| = |rs0| ∧
147 (mk_tape sig (reverse sig rs1@x::ls) (option_hd sig rs2)
149 (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)).
151 axiom sem_copy_strict : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n →
152 copy src dst sig n ⊨ R_copy_strict src dst sig n.
154 lemma sem_unistep : ∀n,l,h.unistep ⊨ R_unistep n l h.
156 @(sem_seq_app ??????? (sem_match_m cfg prg FSUnialpha 2 ???)
157 (sem_seq ?????? (sem_restart_tape ???)
158 (sem_seq ?????? (sem_move_multi ? 2 cfg R ?)
159 (sem_seq ?????? (sem_copy_strict prg cfg FSUnialpha 2 ???)
160 (sem_exec_move …)))))
161 /2 by le_n,sym_not_eq/
162 #ta #tb #HR #state #char #table #Hta_cfg #Hcfg #Hta_prg #Htable
163 #Hbits_obj #Htotaltable
164 #nstate #nchar #m #t #Htuple #Hmatch
165 cases HR -HR #tc * whd in ⊢ (%→?);
166 >Hta_cfg #H cases (H ?? (refl ??)) -H
167 (* prg starts with a bar, so it's not empty *) #_
168 >Hta_prg #H lapply (H ??? (refl ??)) -H *
169 [| cases Htotaltable #ll * #lr #H >H
170 #Hfalse @False_ind cases (Hfalse ll lr) #H1 @H1 //]
171 * #ll * #lr * #Hintable -Htotaltable #Htc
172 * #td * whd in ⊢ (%→?); >Htc
173 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
174 #Htd lapply (Htd ? (refl ??)) -Htd
175 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
176 >(?: list_of_tape ? (mk_tape ? (reverse ? (state@[char])@[bar]) (None ?) [ ]) =
178 [|whd in ⊢ (??%?); >left_mk_tape >reverse_append >reverse_reverse
179 >current_mk_tape >right_mk_tape normalize >append_nil % ]
180 whd in ⊢ (???(???(????%?)??)→?); whd in match (tail ??); #Htd
182 * #te * whd in ⊢ (%→?); >Htd
183 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
184 >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
185 >Htable in Hintable; #Hintable #Hte
187 lapply (cfg_in_table_to_tuple ???? Hcfg ?? Hintable)
188 * #newconfig * #m0 * #lr0 * * #Hlr destruct (Hlr) #Hnewcfg #Hm0
189 cut (∃fo,so.state = fo::so ∧ |so| = n)
190 [ @daemon ] * #fo * #so * #Hstate_exp #Hsolen
191 cut (∃fn,sn,cn.newconfig = fn::sn@[cn] ∧ |sn| = n)
192 [ @daemon ] * #fn * #sn * #cn * #Hnewstate_exp #Hsnlen
193 * #tf * * #_ >Hte >(nth_change_vec ?????? lt_prg)
194 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
195 >Hstate_exp >Hnewstate_exp
196 whd in match (mk_tape ????); whd in match (tape_move ???);
197 #Htf cases (Htf ?????? (refl ??) (refl ??) ?) -Htf
198 [| whd in match (tail ??); >length_append >length_append
199 >Hsolen >length_append >Hsnlen //]
200 #rs1 * #rs2 whd in match (tail ??); * *
202 >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
203 >change_vec_change_vec >append_nil #Htf
206 [@(append_l1_injective ?????? Hrs1rs2) >Hrs1len
207 >length_append >length_append normalize >Hsolen >Hsnlen % ] #Hrs1
208 cut (m0::lr0=rs2) [@(append_l2_injective ?????? Hrs1rs2) >Hrs1 % ] #Hrs2
209 cut (∃ll1. ll@[bar] = bar::ll1)
210 [@daemon (* ll is at the begininng of a table *)] * #ll1 #Hll
211 whd in ⊢ (%→?); >Htf >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
212 >nth_change_vec [2:@leb_true_to_le %]
213 >nth_change_vec [2:@leb_true_to_le %]
214 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
215 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
216 >append_cons <Hrs1 >reverse_append >reverse_single
217 <Hrs2 >(append_cons … bar ll) >Hll >reverse_cons
220 (sem_exec_move ? m0 ?
221 (([cn]@reverse FSUnialpha sn)
222 @fn::reverse FSUnialpha (ll1@(fo::so)@[char])) lr0 … (refl ??) ????)
226 |whd in ⊢ (??%?); >associative_append >associative_append
227 >associative_append %
228 |#Htb >Htb @(eq_vec … (niltape ?)) (* tape by tape *) #i #lei2
229 cases (le_to_or_lt_eq … (le_S_S_to_le …lei2))
230 [#lei1 cases (le_to_or_lt_eq … (le_S_S_to_le …lei1))
231 [#lei0 lapply(le_n_O_to_eq … (le_S_S_to_le …lei0)) #eqi <eqi
233 >nth_change_vec [2:@leb_true_to_le %]
234 (* dimostrare che la tabella e' univoca
235 da cui m = m0 e nchar = cn *)
236 |(* cfg tape *) #eqi >eqi
237 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
238 >nth_change_vec [2:@leb_true_to_le %]
241 |(* prg tape *) #eqi >eqi
242 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
243 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
244 >Hta_prg whd in ⊢ (??%?); @eq_f @(cons_injective_r ? bar bar)
245 >Htable >Hintable >reverse_append >reverse_cons
246 >reverse_reverse >reverse_cons >reverse_reverse
247 >associative_append >associative_append >associative_append
248 >(append_cons ? bar ll) >Hll @eq_f @eq_f <Hstate_exp @eq_f
258 cut ((mk_tape FSUnialpha []
259 (option_hd FSUnialpha
260 (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar])))
262 (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar])))) =
263 midtape ? [ ] bar (fn::sn@[cn;m0]))
264 [cut (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar]) =
266 [>reverse_cons whd in ⊢ (??(??(??%)?)?); >reverse_append >reverse_reverse
267 >append_cons in ⊢ (???%); % ] #Hrev >Hrev % ] #Hmk_tape >Hmk_tape -Hmk_tape
270 >reverse_append >reverse_append
272 <reverse_cons >reverse_cons
275 >(change_vec_commute ????? cfg prg) [2:@eqb_false_to_not_eq %]
276 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
277 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
278 lapply (append_l1_injective ?????? Hrs1rs2)
279 [ >Hsnlen >Hrs1len >length_append >length_append >length_append >length_append
280 normalize >Hsolen >Hsnlen % ]
281 #Hrs1 <Hrs1 >reverse_append #Htg cases (Htg ?? (refl ??)) -Htg
291 match_m cfg prg FSUnialpha 2 ·
292 restart_tape cfg · copy prg cfg FSUnialpha 2 ·
293 cfg_to_obj · tape_move_obj · restart_tape prg · obj_to_cfg.
295 definition tape_map ≝ λA,B:FinSet.λf:A→B.λt.
296 mk_tape B (map ?? f (left ? t))
297 (option_map ?? f (current ? t))
298 (map ?? f (right ? t)).
300 lemma map_list_of_tape: ∀A,B,f,t.
301 list_of_tape B (tape_map ?? f t) = map ?? f (list_of_tape A t).
302 #A #B #f * // normalize // #ls #c #rs <map_append %
305 lemma low_char_current : ∀t.
306 low_char' (current FSUnialpha (tape_map FinBool FSUnialpha bit t))
307 = low_char (current FinBool t).
310 definition low_tapes: ∀M:normalTM.∀c:nconfig (no_states M).Vector ? 3 ≝
311 λM:normalTM.λc:nconfig (no_states M).Vector_of_list ?
312 [tape_map ?? bit (ctape ?? c);
314 ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]);
315 midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
318 lemma obj_low_tapes: ∀M,c.
319 nth obj ? (low_tapes M c) (niltape ?) = tape_map ?? bit (ctape ?? c).
322 lemma cfg_low_tapes: ∀M,c.
323 nth cfg ? (low_tapes M c) (niltape ?) =
325 ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]).
328 lemma prg_low_tapes: ∀M,c.
329 nth prg ? (low_tapes M c) (niltape ?) =
330 midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M))).
333 (* commutation lemma for write *)
334 lemma map_write: ∀t,cout.
335 tape_write ? (tape_map FinBool ? bit t) (char_to_bit_option (low_char cout))
336 = tape_map ?? bit (tape_write ? t cout).
337 #t * // #b whd in match (char_to_bit_option ?);
338 whd in ⊢ (??%%); @eq_f3 [elim t // | // | elim t //]
341 (* commutation lemma for moves *)
342 lemma map_move: ∀t,m.
343 tape_move ? (tape_map FinBool ? bit t) (char_to_move (low_mv m))
344 = tape_map ?? bit (tape_move ? t m).
345 #t * // whd in match (char_to_move ?);
346 [cases t // * // | cases t // #ls #a * //]
349 (* commutation lemma for actions *)
350 lemma map_action: ∀t,cout,m.
351 tape_move ? (tape_write ? (tape_map FinBool ? bit t)
352 (char_to_bit_option (low_char cout))) (char_to_move (low_mv m))
353 = tape_map ?? bit (tape_move ? (tape_write ? t cout) m).
354 #t #cout #m >map_write >map_move %
357 lemma map_move_mono: ∀t,cout,m.
358 tape_move_mono ? (tape_map FinBool ? bit t)
359 〈char_to_bit_option (low_char cout), char_to_move (low_mv m)〉
360 = tape_map ?? bit (tape_move_mono ? t 〈cout,m〉).
364 definition R_unistep_high ≝ λM:normalTM.λt1,t2.
365 ∀c:nconfig (no_states M).
367 t2 = low_tapes M (step ? M c).
369 lemma R_unistep_equiv : ∀M,t1,t2.
370 R_unistep (no_states M) (graph_enum ?? (ntrans M)) (nhalt M) t1 t2 →
371 R_unistep_high M t1 t2.
372 #M #t1 #t2 #H whd whd in match (nconfig ?); #c #Ht1
373 lapply (initial_bar ? (nhalt M) (graph_enum ?? (ntrans M)) (nTM_nog ?)) #Htable
374 (* tup = current tuple *)
375 cut (∃t.t = 〈〈cstate … c,current ? (ctape … c)〉,
376 ntrans M 〈cstate … c,current ? (ctape … c)〉〉) [% //] * #tup #Htup
377 (* tup is in the graph *)
378 cut (mem ? tup (graph_enum ?? (ntrans M)))
379 [@memb_to_mem >Htup @(graph_enum_complete … (ntrans M)) %] #Hingraph
380 (* tupe target = 〈qout,cout,m〉 *)
381 lapply (decomp_target ? (ntrans M 〈cstate … c,current ? (ctape … c)〉))
382 * #qout * #cout * #m #Htg >Htg in Htup; #Htup
384 cut (step FinBool M c = mk_config ?? qout (tape_move ? (tape_write ? (ctape … c) cout) m))
385 [>(config_expand … c) whd in ⊢ (??%?); (* >Htg ?? why not?? *)
386 cut (trans ? M 〈cstate … c, current ? (ctape … c)〉 = 〈qout,cout,m〉) [<Htg %] #Heq1
389 cut (cstate ?? (step FinBool M c) = qout) [>Hstep %] #Hnew_state
391 cut (ctape ?? (step FinBool M c) = tape_move ? (tape_write ? (ctape … c) cout) m)
392 [>Hstep %] #Hnew_tape
393 lapply(H (bits_of_state ? (nhalt M) (cstate ?? c))
394 (low_char (current ? (ctape ?? c)))
395 (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
398 lapply(list_to_table … (nhalt M) …Hingraph) * #ll * #lr #Htable1 %{ll}
399 %{(((bits_of_state ? (nhalt M) qout)@[low_char cout;low_mv m])@lr)}
400 >Htable1 @eq_f <associative_append @eq_f2 // >Htup
401 whd in ⊢ (??%?); @eq_f >associative_append %
402 |>Ht1 >obj_low_tapes >map_list_of_tape elim (list_of_tape ??)
403 [#b @False_ind | #b #tl #Hind #a * [#Ha >Ha //| @Hind]]
406 |%{(bits_of_state ? (nhalt M) (cstate ?? c))} %{(low_char (current ? (ctape ?? c)))}
407 % [% [% [// | cases (current ??) normalize [|#b] % #Hd destruct (Hd)]
408 |>length_map whd in match (length ??); @eq_f //]
410 |>Ht1 >cfg_low_tapes //] -H #H
411 lapply(H (bits_of_state … (nhalt M) qout) (low_char … cout)
412 (low_mv … m) tup ? Hingraph)
413 [>Htup whd in ⊢ (??%?); @eq_f >associative_append %] -H
414 #Ht2 @(eq_vec ? 3 ?? (niltape ?) ?) >Ht2 #i #Hi
415 cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
416 [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
417 [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
419 |>Hi >obj_low_tapes >nth_change_vec //
420 >Ht1 >obj_low_tapes >Hstep @map_action
422 |>Hi >cfg_low_tapes >nth_change_vec_neq
423 [|% whd in ⊢ (??%?→?); #H destruct (H)]
424 >nth_change_vec // >Hnew_state @eq_f @eq_f >Hnew_tape
425 @eq_f2 [|2:%] >Ht1 >obj_low_tapes >map_move_mono >low_char_current %
427 |(* program tapes do not change *)
429 >nth_change_vec_neq [|% whd in ⊢ (??%?→?); #H destruct (H)]
430 >nth_change_vec_neq [|% whd in ⊢ (??%?→?); #H destruct (H)]
431 >Ht1 >prg_low_tapes //