]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/unistep.ma
made executable again
[helm.git] / matita / matita / lib / turing / multi_universal / unistep.ma
1 (*
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.           
5     ||I||                                                            
6     ||T||  
7     ||A||  
8     \   /  This file is distributed under the terms of the       
9      \ /   GNU General Public License Version 2   
10       V_____________________________________________________________*)
11
12
13 include "turing/multi_universal/unistep_aux.ma".
14 include "turing/multi_universal/match.ma".
15
16 definition exec_move ≝ 
17   cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg.
18
19 definition R_exec_move ≝ λt1,t2:Vector (tape FSUnialpha) 3.
20 ∀c,m,ls1,ls2,rs2.
21   nth cfg ? t1 (niltape ?) = mk_tape FSUnialpha (c::ls1@[bar]) (None ?) [ ] → 
22   nth prg ? t1 (niltape ?) = midtape FSUnialpha (ls2@[bar]) m rs2 →
23   only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) →
24   c ≠ bar → m ≠ bar →
25   let new_obj ≝ 
26       tape_move_mono ? (nth obj ? t1 (niltape ?)) 
27         〈char_to_bit_option c, char_to_move m〉 in
28   let next_c ≝ low_char' (current ? new_obj) in 
29   let new_cfg ≝ midtape ? [ ] bar ((reverse ? ls1)@[next_c]) in 
30   let new_prg ≝ midtape FSUnialpha [ ] bar ((reverse ? ls2)@m::rs2) in
31   t2 = Vector_of_list ? [new_obj;new_cfg;new_prg].
32
33
34 lemma sem_exec_move: exec_move ⊨ R_exec_move.
35 @(sem_seq_app ??????? sem_cfg_to_obj1
36   (sem_seq ?????? sem_tape_move_obj
37    (sem_seq ?????? (sem_restart_tape ???) sem_obj_to_cfg1))) //
38 #ta #tout * #t1 * #semM1 * #t2 * #semM2 * #t3 * #semM3 #semM4 
39 #c #m #ls1 #ls2 #rs2 #Hcfg #Hprg #Honlybits #Hc #Hm 
40 (* M1 = cfg_to_obj *)
41 lapply (semM1 … Hcfg Hc) #Ht1 
42 (* M2 = tape_move *)
43 whd in semM2; >Ht1 in semM2; -Ht1
44 >nth_change_vec_neq [2:@eqb_false_to_not_eq %] 
45 >nth_change_vec_neq [2:@eqb_false_to_not_eq %] 
46 >Hprg #Ht2 lapply (Ht2 … (refl ??)) -Ht2
47 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
48 >nth_change_vec // >change_vec_commute [2:@eqb_false_to_not_eq %]
49 >change_vec_change_vec #Ht2 
50 (* M3 = restart prg *) 
51 whd in semM3; >Ht2 in semM3; #semM3 lapply (semM3 … (refl ??)); -semM3
52 >nth_change_vec_neq [2:@eqb_false_to_not_eq %] 
53 >nth_change_vec_neq [2:@eqb_false_to_not_eq %] #Ht3
54 (* M4 = obj_to_cfg *) 
55 whd in semM4; >Ht3 in semM4; 
56 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
57 >nth_change_vec [2:@leb_true_to_le %] #semM4 lapply (semM4 … (refl ??)) -semM4
58 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
59 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
60 >nth_change_vec [2:@leb_true_to_le %] #semM4 >(semM4 ?)
61   [(* tape by tape *)
62    @(eq_vec … (niltape ?)) #i #lei2 
63    cases (le_to_or_lt_eq … (le_S_S_to_le …lei2))
64     [#lei1 cases (le_to_or_lt_eq … (le_S_S_to_le …lei1))
65       [#lei0 lapply(le_n_O_to_eq … (le_S_S_to_le …lei0)) #eqi <eqi
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_neq [2:@eqb_false_to_not_eq %]
69        >nth_change_vec [2:@leb_true_to_le %] %
70       |#Hi >Hi (* obj tape *)
71        >nth_change_vec [2:@leb_true_to_le %] whd in ⊢ (???%);
72        >reverse_cons >reverse_append >reverse_single 
73        whd in match (option_hd ??); whd in match (tail ??);
74        whd in ⊢ (??%?); %
75       ]
76     |#Hi >Hi (* prg tape *)
77      >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
78      >nth_change_vec [2:@leb_true_to_le %] whd in ⊢ (???%);
79      >Hprg whd in match (list_of_tape ??); >reverse_append
80      >reverse_single % 
81     ]
82   |#b #Hcurrent @(list_of_tape_write ? is_bit … (char_to_bit_option c) ? Honlybits)
83     [#b0 cases c
84       [#b1 whd in ⊢ (??%?→?); #Hbit destruct (Hbit) %
85       |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
86       |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
87       ]
88     |>(list_of_tape_move … (char_to_move m)) @current_in_list @Hcurrent
89     ]
90   ]
91 qed.
92        
93 definition unistep ≝ 
94   match_m cfg prg FSUnialpha 2 · 
95   restart_tape cfg 2 · mmove cfg ? 2 R · copy prg cfg FSUnialpha 2 ·
96   exec_move.
97
98 (*
99 definition legal_tape ≝ λn,l,h,t.
100   ∃state,char,table.
101   nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) →
102   is_config n (bar::state@[char]) →  
103   nth prg ? t1 (niltape ?) = midtape ? [ ] bar table →
104   bar::table = table_TM n l h → *)
105   
106 definition deterministic_tuples ≝ λn,h,l.
107  ∀t1,t2,conf,out1,out2.
108   is_config n conf →
109   mem ? t1 l → mem ? t2 l →
110   tuple_encoding n h t1 = conf@out1 →
111     tuple_encoding n h t2 = conf@out2 → out1 = out2.
112
113 definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3.
114   ∀state,char,table.
115   (* cfg *)
116   nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) →
117   is_config n (bar::state@[char]) →  
118   (* prg *)
119   nth prg ? t1 (niltape ?) = midtape ? [ ] bar table →
120   bar::table = table_TM n l h →
121   (* obj *)
122   only_bits (list_of_tape ? (nth obj ? t1 (niltape ?))) →
123   (* deterministic tuples *)
124   deterministic_tuples n h l →
125   let conf ≝ (bar::state@[char]) in
126   (∃ll,lr.bar::table = ll@conf@lr) →
127 (*
128     ∃nstate,nchar,m,t. tuple_encoding n h t = (conf@nstate@[nchar;m]) ∧ 
129     mem ? t l ∧  *)
130     ∀nstate,nchar,m,t. 
131     tuple_encoding n h t = (conf@nstate@[nchar;m])→ 
132     mem ? t l →
133     let new_obj ≝ 
134      tape_move_mono ? (nth obj ? t1 (niltape ?)) 
135        〈char_to_bit_option nchar,char_to_move m〉 in
136     let next_char ≝ low_char' (current ? new_obj) in
137     t2 = 
138       change_vec ??
139         (change_vec ?? t1 (midtape ? [ ] bar (nstate@[next_char])) cfg)
140         new_obj obj.
141         
142 lemma lt_obj : obj < 3. // qed.
143 lemma lt_cfg : cfg < 3. // qed.
144 lemma lt_prg : prg < 3. // qed.
145
146 definition R_copy_strict ≝ 
147   λsrc,dst,sig,n.λint,outt: Vector (tape sig) (S n).
148   ((current ? (nth src ? int (niltape ?)) = None ? ∨
149     current ? (nth dst ? int (niltape ?)) = None ?) → outt = int) ∧
150   (∀ls,x,x0,rs,ls0,rs0. 
151     nth src ? int (niltape ?) = midtape sig ls x rs →
152     nth dst ? int (niltape ?) = midtape sig ls0 x0 rs0 →
153     |rs0| ≤ |rs| → 
154     (∃rs1,rs2.rs = rs1@rs2 ∧ |rs1| = |rs0| ∧
155      outt = change_vec ?? 
156             (change_vec ?? int  
157               (mk_tape sig (reverse sig rs1@x::ls) (option_hd sig rs2)
158             (tail sig rs2)) src)
159             (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)).
160
161 lemma sem_copy_strict : ∀src,dst,sig,n. src ≠ dst → src < S n → dst < S n → 
162   copy src dst sig n ⊨ R_copy_strict src dst sig n.
163 #src #dst #sig #n #Hneq #Hsrc #Hdst @(Realize_to_Realize … (sem_copy …)) //
164 #ta #tb * #Htb1 #Htb2 % [ @Htb1 ]
165 #ls #x #x0 #rs #ls0 #rs0 #Htamid_src #Htamid_dst #Hlen
166 cases (Htb2 … Htamid_src Htamid_dst) -Htb1 -Htb2
167 [ * #rs1 * #rs2 * * #Hrs0 #Heq #Htb <Heq in Hlen; >Hrs0 >length_append
168   >(plus_n_O (|rs1|)) #Hlen cut (|rs2| ≤ 0) [ /2 by le_plus_to_le/ ]
169   #Hlenrs2 cut (rs2 = [ ]) 
170   [ cases rs2 in Hlenrs2; // #r3 #rs3 normalize #H @False_ind cases (not_le_Sn_O (|rs3|)) /2/ ]
171   #Hrs2 destruct (Hrs2) >append_nil in Hrs0; #Hrs0 destruct (Hrs0) -Hlenrs2 -Hlen
172   <plus_n_O <plus_n_O %{rs} %{[ ]} >append_nil % [ % // ] @Htb
173 | * #rs1 * #rs2 #H %{rs1} %{rs2} @H ]
174 qed.
175
176 lemma config_to_len : ∀n,b,q,c.is_config n (b::q@[c]) → |q| = S n.
177 #n #b #q #c * #q0 * #cin * * * #_ #_ #Hq0 #H >(?:q = q0) //
178 lapply (cons_injective_r ????? H) #H1 @(append_l1_injective … H1)
179 lapply (eq_f … (length ?) … H) normalize >length_append >length_append
180 <plus_n_Sm <plus_n_Sm <plus_n_O <plus_n_O #Hlen destruct (Hlen) //
181 qed.
182
183 lemma sem_unistep_low : ∀n,l,h.unistep ⊨ R_unistep n l h.
184 #n #l #h
185 @(sem_seq_app ??????? (sem_match_m cfg prg FSUnialpha 2 ???)
186   (sem_seq ?????? (sem_restart_tape ???)
187    (sem_seq ?????? (sem_move_multi ? 2 cfg R ?)
188     (sem_seq ?????? (sem_copy_strict prg cfg FSUnialpha 2 ???)
189      (sem_exec_move …)))))
190   /2 by le_n,sym_not_eq/
191 #ta #tb #HR #state #char #table #Hta_cfg #Hcfg #Hta_prg #Htable
192 #Hbits_obj #Hdeterm #Hmatching
193 #nstate #nchar #m #t #Htuple #Hmatch
194 cases HR -HR #tc * whd in ⊢ (%→?); 
195 >Hta_cfg #H cases (H ?? (refl ??)) -H 
196 (* prg starts with a bar, so it's not empty *) #_
197 >Hta_prg #H lapply (H ??? (refl ??)) -H *
198 [| cases Hmatching #ll * #lr #H >H 
199    #Hfalse @False_ind cases (Hfalse ll lr) #H1 @H1 //]
200 * #ll * #lr * #Hintable  #Htc
201 * #td * whd in ⊢ (%→?); >Htc
202 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
203 #Htd lapply (Htd ? (refl ??)) -Htd
204 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
205 >(?: list_of_tape ? (mk_tape ? (reverse ? (state@[char])@[bar]) (None ?) [ ]) =
206      bar::state@[char]) 
207 [|whd in ⊢ (??%?); >left_mk_tape >reverse_append >reverse_reverse
208   >current_mk_tape >right_mk_tape [| #_ %2 % ] normalize >append_nil % ]
209 whd in ⊢ (???(???(????%?)??)→?); whd in match (tail ??); #Htd
210 (* move cfg to R *)
211 * #te * whd in ⊢ (%→?); >Htd
212 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
213 >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec [2:@leb_true_to_le %]
214 cut (∃ll1.ll@[bar] = bar::ll1) 
215 [ cases ll in Hintable; [ #_ %{[ ]} % ] 
216   #llhd #lltl normalize #H destruct (H) %{(lltl@[bar])} % ] * #ll1 #Hll1
217 >Htable in Hintable; #Hintable #Hte
218 (* copy *)
219 lapply (table_to_list ???? Hcfg ?? Hintable) * #out * #lr0 * #t0
220 * * #Hlr #Htuple_t0 #mem_t0 
221 cut (out = nstate@[nchar;m])
222   [@(Hdeterm … Hcfg mem_t0 Hmatch Htuple_t0 Htuple)] #Hout
223 >(append_cons ? nchar) in Htuple; #Htuple
224 lapply (tuple_to_config ?????? Hcfg … Htuple) #newconfig
225 cut (∃fo,so.state = fo::so ∧ |so| = n)
226 [ lapply (config_to_len … Hcfg) cases state [ normalize #H destruct (H) ]
227   #fo #so normalize #H destruct (H) %{fo} %{so} % // ]
228 * #fo * #so * #Hstate_exp #Hsolen
229 cut (∃fn,sn.nstate = fn::sn  ∧ |sn| = n)
230 [ lapply (config_to_len … newconfig) cases nstate [ normalize #H destruct (H) ]
231   #fn #sn normalize #H destruct (H) %{fn} %{sn} % // ]
232 * #fn * #sn * #Hnewstate_exp #Hsnlen
233 >Hstate_exp >Hout in Hlr; >Hnewstate_exp whd in ⊢ (???%→?); #Hlr
234 * #tf * * #_ >Hte >(nth_change_vec ?????? lt_prg)
235 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
236 >Hstate_exp >Hnewstate_exp >Hlr
237 whd in match (mk_tape ????); whd in match (tape_move ???);
238 #Htf cases (Htf ?????? (refl ??) (refl ??) ?) -Htf
239 [| whd in match (tail ??); >length_append >length_append 
240    >Hsolen >length_append >Hsnlen normalize //]
241 #rs1 * #rs2 whd in match (tail ??); * *
242 #Hrs1rs2 #Hrs1len
243 >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
244 >change_vec_change_vec >append_nil #Htf 
245 (* exec_move *)
246 >append_cons in Hrs1rs2; >associative_append #Hrs1rs2
247 cut ((sn@[nchar])=rs1)
248   [@(append_l1_injective ?????? Hrs1rs2) >Hrs1len 
249    >length_append >length_append normalize >Hsolen >Hsnlen % ] #Hrs1  
250 cut (m::lr0=rs2) [@(append_l2_injective ?????? Hrs1rs2) >Hrs1 % ] #Hrs2 
251 whd in ⊢ (%→?); >Htf >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
252 >nth_change_vec [2:@leb_true_to_le %] 
253 >nth_change_vec [2:@leb_true_to_le %]
254 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
255 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
256 >append_cons <Hrs1 >reverse_append >reverse_single 
257 <Hrs2 >(append_cons … bar ll) >Hll1 >reverse_cons
258 #sem_exec_move 
259 lapply
260   (sem_exec_move ? m ?
261     (([nchar]@reverse FSUnialpha sn)
262           @fn::reverse FSUnialpha (ll1@(fo::so)@[char])) lr0 … (refl ??) ????) 
263   [ cut (tuple_TM (S n) (tuple_encoding n h t)) // >Htuple
264     * #qin * #cin * #qout * #cout * #mv * * * * #_ #Hmv #_ #_
265     normalize >(?: bar::qin@cin::qout@[cout;mv] = (bar::qin@cin::qout@[cout])@[mv])
266     [| normalize >associative_append normalize >associative_append % ] 
267     >(?: bar::(state@[char])@(nstate@[nchar])@[m] = (bar::(state@[char])@(nstate@[nchar]))@[m]) 
268     [|normalize >associative_append >associative_append >associative_append >associative_append >associative_append % ] 
269     #Heq lapply (append_l2_injective_r ?????? Heq) // #H destruct (H) //
270   | cases newconfig #qout * #cout * * * #_ #Hcout #_ #H destruct (H) -H
271     lapply (append_l2_injective_r ?????? e0) // #H destruct (H) @Hcout 
272   |@Hbits_obj
273   |whd in ⊢ (??%?); >associative_append >associative_append 
274    >associative_append  %
275   |#Htb >Htb @(eq_vec … (niltape ?)) (* tape by tape *) #i #lei2 
276    cases (le_to_or_lt_eq … (le_S_S_to_le …lei2))
277     [#lei1 cases (le_to_or_lt_eq … (le_S_S_to_le …lei1))
278       [#lei0 lapply(le_n_O_to_eq … (le_S_S_to_le …lei0)) #eqi <eqi
279        (* obj tape *)
280        >nth_change_vec [2:@leb_true_to_le %] % 
281       |(* cfg tape *) #eqi >eqi
282        >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
283        >nth_change_vec [2:@leb_true_to_le %]  
284        whd in ⊢ (??%?); cut (∀A.∀l:list A.[]@l = l) [//] #Hnil >Hnil
285        >reverse_append >reverse_single >reverse_reverse %
286        (* idem *)
287       ]
288     |(* prg tape *) #eqi >eqi
289       >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
290       >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
291       >Hta_prg whd in ⊢ (??%?); @eq_f @(cons_injective_r ? bar bar)
292       >Htable >Hintable >reverse_append >reverse_cons
293       >reverse_reverse >reverse_cons >reverse_reverse
294       >associative_append >associative_append >associative_append
295       >(append_cons ? bar ll) >Hll1 @eq_f @eq_f <Hstate_exp @eq_f
296       >Hnewstate_exp >Hlr normalize >associative_append >associative_append %
297     ]
298   ]
299 qed.
300
301 definition tape_map ≝ λA,B:FinSet.λf:A→B.λt.
302   mk_tape B (map ?? f (left ? t)) 
303     (option_map ?? f (current ? t)) 
304     (map ?? f (right ? t)).
305
306 (* da spostare *)
307 lemma map_reverse: ∀A,B,f,l. 
308   map ?? f (reverse A l) = reverse B (map ?? f l).
309 #A #B #f #l elim l //
310 #a #l1 #Hind >reverse_cons >reverse_cons <map_append @eq_f2 //
311 qed.
312
313 lemma map_list_of_tape: ∀A,B,f,t.
314   list_of_tape B (tape_map ?? f t) = map ?? f (list_of_tape A t).
315 #A #B #f * // 
316   [#a #l normalize >rev_append_def >rev_append_def >append_nil
317    >append_nil <map_append <map_reverse @eq_f2 //
318   |#rs #a #ls normalize >rev_append_def >rev_append_def
319    >append_nil >append_nil <map_append normalize 
320    @eq_f2 //
321   ]
322 qed.
323
324 lemma low_char_current : ∀t.
325   low_char' (current FSUnialpha (tape_map FinBool FSUnialpha bit t))
326   = low_char (current FinBool t).
327 * // qed.
328
329 definition low_tapes: ∀M:normalTM.∀c:nconfig (no_states M).Vector ? 3 ≝ 
330 λM:normalTM.λc:nconfig (no_states M).Vector_of_list ?
331   [tape_map ?? bit (ctape ?? c);
332    midtape ? [ ] bar 
333     ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]);
334    midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
335   ].
336
337 lemma obj_low_tapes: ∀M,c.
338   nth obj ? (low_tapes M c) (niltape ?) = tape_map ?? bit (ctape ?? c).
339 // qed.
340
341 lemma cfg_low_tapes: ∀M,c.
342   nth cfg ? (low_tapes M c) (niltape ?) = 
343   midtape ? [ ] bar 
344     ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]).
345 // qed.
346
347 lemma prg_low_tapes: ∀M,c.
348   nth prg ? (low_tapes M c) (niltape ?) = 
349   midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M))).
350 // qed.
351
352 (* commutation lemma for write *)
353 lemma map_write: ∀t,cout.
354  tape_write ? (tape_map FinBool ? bit t) (char_to_bit_option (low_char cout))
355   = tape_map ?? bit (tape_write ? t cout).
356 #t * // #b whd in match (char_to_bit_option ?);
357 whd in ⊢ (??%%); @eq_f3 [elim t // | // | elim t //]
358 qed.
359
360 (* commutation lemma for moves *)
361 lemma map_move: ∀t,m.
362  tape_move ? (tape_map FinBool ? bit t) (char_to_move (low_mv m))
363   = tape_map ?? bit (tape_move ? t m).
364 #t * // whd in match (char_to_move ?);
365   [cases t // * // | cases t // #ls #a * //]
366 qed.
367   
368 (* commutation lemma for actions *)
369 lemma map_action: ∀t,cout,m.
370  tape_move ? (tape_write ? (tape_map FinBool ? bit t)
371     (char_to_bit_option (low_char cout))) (char_to_move (low_mv m)) 
372  = tape_map ?? bit (tape_move ? (tape_write ? t cout) m).
373 #t #cout #m >map_write >map_move % 
374 qed. 
375
376 lemma map_move_mono: ∀t,cout,m.
377  tape_move_mono ? (tape_map FinBool ? bit t)
378   〈char_to_bit_option (low_char cout), char_to_move (low_mv m)〉
379  = tape_map ?? bit (tape_move_mono ? t 〈cout,m〉).
380 @map_action
381 qed. 
382
383 definition R_unistep_high ≝ λM:normalTM.λt1,t2.
384 ∀c:nconfig (no_states M).
385   t1 = low_tapes M c → 
386   t2 = low_tapes M (step ? M c). 
387
388 lemma R_unistep_equiv : ∀M,t1,t2. 
389   R_unistep (no_states M) (graph_enum ?? (ntrans M)) (nhalt M) t1 t2 →
390   R_unistep_high M t1 t2.
391 #M #t1 #t2 #H whd whd in match (nconfig ?); #c #Ht1
392 lapply (initial_bar ? (nhalt M) (graph_enum ?? (ntrans M)) (nTM_nog ?)) #Htable
393 (* tup = current tuple *)
394 cut (∃t.t = 〈〈cstate … c,current ? (ctape … c)〉, 
395              ntrans M 〈cstate … c,current ? (ctape … c)〉〉) [% //] * #tup #Htup
396 (* tup is in the graph *)
397 cut (mem ? tup (graph_enum ?? (ntrans M)))
398   [@memb_to_mem >Htup @(graph_enum_complete … (ntrans M)) %] #Hingraph
399 (* tupe target = 〈qout,cout,m〉 *)
400 lapply (decomp_target ? (ntrans M 〈cstate … c,current ? (ctape … c)〉))
401 * #qout * #cout * #m #Htg >Htg in Htup; #Htup
402 (* new config *)
403 cut (step FinBool M c = mk_config ?? qout (tape_move ? (tape_write ? (ctape … c) cout) m))
404   [>(config_expand … c) whd in ⊢ (??%?); (* >Htg ?? why not?? *)
405    cut (trans ? M 〈cstate  … c, current ? (ctape … c)〉 = 〈qout,cout,m〉) [<Htg %] #Heq1 
406    >Heq1 %] #Hstep
407 (* new state *)
408 cut (cstate ?? (step FinBool M c) = qout) [>Hstep %] #Hnew_state
409 (* new tape *)
410 cut (ctape ?? (step FinBool M c) = tape_move ? (tape_write ? (ctape … c) cout) m)
411   [>Hstep %] #Hnew_tape
412 lapply(H (bits_of_state ? (nhalt M) (cstate ?? c)) 
413          (low_char (current ? (ctape ?? c)))
414          (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
415          ???????)        
416 [<Htable
417  lapply(list_to_table … (nhalt M) …Hingraph) * #ll * #lr #Htable1 %{ll} 
418  %{(((bits_of_state ? (nhalt M) qout)@[low_char cout;low_mv m])@lr)} 
419  >Htable1 @eq_f <associative_append @eq_f2 // >Htup
420  whd in ⊢ (??%?); @eq_f >associative_append %
421 |#tx #ty #conf #outx #outy #isconf #memx #memy #tuplex #tupley
422  @(deterministic M … (refl ??) memx memy isconf tuplex tupley)
423 |>Ht1 >obj_low_tapes >map_list_of_tape elim (list_of_tape ??) 
424   [#b @False_ind | #b #tl #Hind #a * [#Ha >Ha //| @Hind]]
425 |@sym_eq @Htable
426 |>Ht1 %
427 |%{(bits_of_state ? (nhalt M) (cstate ?? c))} %{(low_char (current ? (ctape ?? c)))}
428  % [% [% [// | cases (current ??) normalize [|#b] % #Hd destruct (Hd)]
429       |>length_map whd in match (length ??); @eq_f //]
430    |//]
431 |>Ht1 >cfg_low_tapes //] -H #H 
432 lapply(H (bits_of_state … (nhalt M) qout) (low_char … cout) 
433          (low_mv … m) tup ? Hingraph)
434   [>Htup whd in ⊢ (??%?); @eq_f >associative_append %] -H
435 #Ht2 @(eq_vec ? 3 ?? (niltape ?) ?) >Ht2 #i #Hi 
436 cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
437   [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
438     [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
439       [@False_ind /2/
440       |>Hi >obj_low_tapes >nth_change_vec //
441        >Ht1 >obj_low_tapes >Hstep @map_action 
442       ]
443     |>Hi >cfg_low_tapes >nth_change_vec_neq 
444       [|% whd in ⊢ (??%?→?);  #H destruct (H)]
445      >nth_change_vec // >Hnew_state @eq_f @eq_f >Hnew_tape 
446      @eq_f2 [|2:%] >Ht1 >obj_low_tapes >map_move_mono >low_char_current %
447     ]
448   |(* program tapes do not change *)
449    >Hi >prg_low_tapes 
450    >nth_change_vec_neq [|% whd in ⊢ (??%?→?);  #H destruct (H)]
451    >nth_change_vec_neq [|% whd in ⊢ (??%?→?);  #H destruct (H)]
452    >Ht1 >prg_low_tapes //
453   ]
454 qed.
455
456 lemma sem_unistep : ∀M.unistep ⊨ R_unistep_high M.
457 #M @(Realize_to_Realize … (R_unistep_equiv …)) //
458 qed.