]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/unistep.ma
sem_exec_move completed
[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
15 definition exec_move ≝ 
16   cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg.
17
18 definition R_exec_move ≝ λt1,t2:Vector (tape FSUnialpha) 3.
19 ∀c,m,ls1,ls2,rs2.
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 ?))) →
23   c ≠ bar → m ≠ bar →
24   let new_obj ≝ 
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].
31
32
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 
39 (* M1 = cfg_to_obj *)
40 lapply (semM1 … Hcfg Hc) #Ht1 
41 (* M2 = tape_move *)
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
53 (* M4 = obj_to_cfg *) 
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 ?)
60   [(* tape by tape *)
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 ??);
73        whd in ⊢ (??%?); %
74       ]
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
79      >reverse_single % 
80     ]
81   |#b #Hcurrent @(list_of_tape_write ? is_bit … (char_to_bit_option c) ? Honlybits)
82     [#b0 cases c
83       [#b1 whd in ⊢ (??%?→?); #Hbit destruct (Hbit) %
84       |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
85       |whd in ⊢ (??%?→?); #Hbit destruct (Hbit)
86       ]
87     |>(list_of_tape_move … (char_to_move m)) @current_in_list @Hcurrent
88     ]
89   ]
90 qed.
91        
92 definition unistep ≝ 
93   match_m cfg prg FSUnialpha 2 · 
94   restart_tape cfg 2 · mmove cfg ? 2 R · copy prg cfg FSUnialpha 2 ·
95   cfg_to_obj · tape_move_obj · restart_tape prg 2 · obj_to_cfg.
96
97 (*
98 definition legal_tape ≝ λn,l,h,t.
99   ∃state,char,table.
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 → *)
104
105 definition R_unistep ≝ λn,l,h.λt1,t2: Vector ? 3.
106   ∀state,char,table.
107   (* cfg *)
108   nth cfg ? t1 (niltape ?) = midtape ? [ ] bar (state@[char]) →
109   is_config n (bar::state@[char]) →  
110   (* prg *)
111   nth prg ? t1 (niltape ?) = midtape ? [ ] bar table →
112   bar::table = table_TM n l h →
113   (* obj *)
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) →
117 (*
118     ∃nstate,nchar,m,t. tuple_encoding n h t = (conf@nstate@[nchar;m]) ∧ 
119     mem ? t l ∧  *)
120     ∀nstate,nchar,m,t. 
121     tuple_encoding n h t = (conf@nstate@[nchar;m])→ 
122     mem ? t l →
123     let new_obj ≝ 
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
127     t2 = 
128       change_vec ??
129         (change_vec ?? t1 (midtape ? [ ] bar (nstate@[next_char])) cfg)
130         new_obj obj.
131         
132 lemma lt_obj : obj < 3. // qed.
133 lemma lt_cfg : cfg < 3. // qed.
134 lemma lt_prg : prg < 3. // qed.
135
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 →
143     |rs0| ≤ |rs| → 
144     (∃rs1,rs2.rs = rs1@rs2 ∧ |rs1| = |rs0| ∧
145      outt = change_vec ?? 
146             (change_vec ?? int  
147               (mk_tape sig (reverse sig rs1@x::ls) (option_hd sig rs2)
148             (tail sig rs2)) src)
149             (mk_tape sig (reverse sig rs1@x::ls0) (None sig) []) dst)).
150
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.
153
154 lemma sem_unistep : ∀n,l,h.unistep ⊨ R_unistep n l h.
155 #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_seq ?????? sem_cfg_to_obj1
161       (sem_seq ?????? sem_tape_move_obj
162        (sem_seq ?????? (sem_restart_tape ???) sem_obj_to_cfg)))))))
163   /2 by le_n,sym_not_eq/
164 #ta #tb #HR #state #char #table #Hta_cfg #Hcfg #Hta_prg #Htable
165 #Hbits_obj #Htotaltable
166 #nstate #nchar #m #t #Htuple #Hmatch
167 cases HR -HR #tc * whd in ⊢ (%→?); 
168 >Hta_cfg #H cases (H ?? (refl ??)) -H 
169 (* prg starts with a bar, so it's not empty *) #_
170 >Hta_prg #H lapply (H ??? (refl ??)) -H *
171 [| cases Htotaltable #ll * #lr #H >H 
172    #Hfalse @False_ind cases (Hfalse ll lr) #H1 @H1 //]
173 * #ll * #lr * #Hintable -Htotaltable #Htc
174 * #td * whd in ⊢ (%→?); >Htc
175 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
176 #Htd lapply (Htd ? (refl ??)) -Htd
177 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
178 >(?: list_of_tape ? (mk_tape ? (reverse ? (state@[char])@[bar]) (None ?) [ ]) =
179      bar::state@[char]) 
180 [|whd in ⊢ (??%?); >left_mk_tape >reverse_append >reverse_reverse
181   >current_mk_tape >right_mk_tape normalize >append_nil % ]
182 whd in ⊢ (???(???(????%?)??)→?); whd in match (tail ??); #Htd
183 (* move cfg to R *)
184 * #te * whd in ⊢ (%→?); >Htd
185 >change_vec_commute [|@sym_not_eq //] >change_vec_change_vec
186 >nth_change_vec_neq [|@sym_not_eq //] >nth_change_vec //
187 >Htable in Hintable; #Hintable #Hte
188 (* copy *)
189 cases (cfg_in_table_to_tuple ???? Hcfg ?? Hintable)
190 #newstate * #m0 * #lr0 * * #Hlr destruct (Hlr) #Hnewcfg #Hm0
191 cut (∃fo,so,co.state = fo::so@[co] ∧ |so| = n)
192 [ @daemon ] * #fo * #so * #co * #Hstate_exp #Hsolen
193 cut (∃fn,sn,cn.newstate = fn::sn@[cn] ∧ |sn| = n)
194 [ @daemon ] * #fn * #sn * #cn * #Hnewstate_exp #Hsnlen
195 * #tf * * #_ >Hte >(nth_change_vec ?????? lt_prg)
196 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
197 >Hstate_exp >Hnewstate_exp
198 whd in match (mk_tape ????); whd in match (tape_move ???);
199 #Htf cases (Htf ?????? (refl ??) (refl ??) ?) -Htf
200 [| whd in match (tail ??); >length_append >length_append 
201    >Hsolen >length_append >length_append >Hsnlen 
202    <plus_n_Sm <plus_n_Sm <plus_n_Sm <plus_n_O <plus_n_O normalize // ]
203 #rs1 * #rs2 whd in match (tail ??); * *
204 >append_cons #Hrs1rs2 #Hrs1len
205 >change_vec_change_vec >change_vec_commute [|@sym_not_eq //]
206 >change_vec_change_vec #Htf 
207 (* cfg to obj *)
208 * #tg * whd in ⊢ (%→?); >Htf
209 >nth_change_vec_neq [|@sym_not_eq //] >(nth_change_vec ?????? lt_cfg)
210 lapply (append_l1_injective ?????? Hrs1rs2)
211 [ >Hsnlen >Hrs1len >length_append >length_append >length_append >length_append
212   normalize >Hsolen >Hsnlen % ] #Hrs1 <Hrs1 >reverse_append >reverse_single
213   >associative_append #Htg lapply (Htg … (refl ??) Hm0) -Htg
214   (* simplifying tg *)
215   >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
216   >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
217
218
219
220
221
222   cut ((mk_tape FSUnialpha []
223      (option_hd FSUnialpha
224       (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar])))
225      (tail FSUnialpha
226       (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar])))) =
227       midtape ? [ ] bar (fn::sn@[cn;m0]))
228   [cut (reverse FSUnialpha (m0::[]@reverse FSUnialpha (sn@[cn])@[fn; bar]) =
229           bar::fn::sn@[cn;m0])
230     [>reverse_cons whd in ⊢ (??(??(??%)?)?); >reverse_append >reverse_reverse
231      >append_cons in ⊢ (???%); % ] #Hrev >Hrev % ] #Hmk_tape >Hmk_tape -Hmk_tape
232    >change_vec_commute
233
234   >reverse_append >reverse_append
235   check reverse_cons
236   <reverse_cons >reverse_cons
237    -Htg #Htg 
238
239 >(change_vec_commute ????? cfg prg) [2:@eqb_false_to_not_eq %]
240 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
241 >nth_change_vec_neq [2:@eqb_false_to_not_eq %]
242 lapply (append_l1_injective ?????? Hrs1rs2)
243 [ >Hsnlen >Hrs1len >length_append >length_append >length_append >length_append
244   normalize >Hsolen >Hsnlen % ]
245 #Hrs1 <Hrs1 >reverse_append #Htg cases (Htg ?? (refl ??)) -Htg
246 cases m0
247   [#mv #_ #Htg #_
248
249    
250       
251       
252
253 [ * 
254  
255   match_m cfg prg FSUnialpha 2 ·
256   restart_tape cfg · copy prg cfg FSUnialpha 2 ·
257   cfg_to_obj · tape_move_obj · restart_tape prg · obj_to_cfg.
258
259 definition tape_map ≝ λA,B:FinSet.λf:A→B.λt.
260   mk_tape B (map ?? f (left ? t)) 
261     (option_map ?? f (current ? t)) 
262     (map ?? f (right ? t)).
263     
264 lemma map_list_of_tape: ∀A,B,f,t.
265   list_of_tape B (tape_map ?? f t) = map ?? f (list_of_tape A t).
266 #A #B #f * // normalize // #ls #c #rs <map_append %
267 qed.
268
269 lemma low_char_current : ∀t.
270   low_char' (current FSUnialpha (tape_map FinBool FSUnialpha bit t))
271   = low_char (current FinBool t).
272 * // qed.
273
274 definition low_tapes: ∀M:normalTM.∀c:nconfig (no_states M).Vector ? 3 ≝ 
275 λM:normalTM.λc:nconfig (no_states M).Vector_of_list ?
276   [tape_map ?? bit (ctape ?? c);
277    midtape ? [ ] bar 
278     ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]);
279    midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
280   ].
281
282 lemma obj_low_tapes: ∀M,c.
283   nth obj ? (low_tapes M c) (niltape ?) = tape_map ?? bit (ctape ?? c).
284 // qed.
285
286 lemma cfg_low_tapes: ∀M,c.
287   nth cfg ? (low_tapes M c) (niltape ?) = 
288   midtape ? [ ] bar 
289     ((bits_of_state ? (nhalt M) (cstate ?? c))@[low_char (current ? (ctape ?? c))]).
290 // qed.
291
292 lemma prg_low_tapes: ∀M,c.
293   nth prg ? (low_tapes M c) (niltape ?) = 
294   midtape ? [ ] bar (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M))).
295 // qed.
296
297 (* commutation lemma for write *)
298 lemma map_write: ∀t,cout.
299  tape_write ? (tape_map FinBool ? bit t) (char_to_bit_option (low_char cout))
300   = tape_map ?? bit (tape_write ? t cout).
301 #t * // #b whd in match (char_to_bit_option ?);
302 whd in ⊢ (??%%); @eq_f3 [elim t // | // | elim t //]
303 qed.
304
305 (* commutation lemma for moves *)
306 lemma map_move: ∀t,m.
307  tape_move ? (tape_map FinBool ? bit t) (char_to_move (low_mv m))
308   = tape_map ?? bit (tape_move ? t m).
309 #t * // whd in match (char_to_move ?);
310   [cases t // * // | cases t // #ls #a * //]
311 qed.
312   
313 (* commutation lemma for actions *)
314 lemma map_action: ∀t,cout,m.
315  tape_move ? (tape_write ? (tape_map FinBool ? bit t)
316     (char_to_bit_option (low_char cout))) (char_to_move (low_mv m)) 
317  = tape_map ?? bit (tape_move ? (tape_write ? t cout) m).
318 #t #cout #m >map_write >map_move % 
319 qed. 
320
321 lemma map_move_mono: ∀t,cout,m.
322  tape_move_mono ? (tape_map FinBool ? bit t)
323   〈char_to_bit_option (low_char cout), char_to_move (low_mv m)〉
324  = tape_map ?? bit (tape_move_mono ? t 〈cout,m〉).
325 @map_action
326 qed. 
327
328 definition R_unistep_high ≝ λM:normalTM.λt1,t2.
329 ∀c:nconfig (no_states M).
330   t1 = low_tapes M c → 
331   t2 = low_tapes M (step ? M c). 
332
333 lemma R_unistep_equiv : ∀M,t1,t2. 
334   R_unistep (no_states M) (graph_enum ?? (ntrans M)) (nhalt M) t1 t2 →
335   R_unistep_high M t1 t2.
336 #M #t1 #t2 #H whd whd in match (nconfig ?); #c #Ht1
337 lapply (initial_bar ? (nhalt M) (graph_enum ?? (ntrans M)) (nTM_nog ?)) #Htable
338 (* tup = current tuple *)
339 cut (∃t.t = 〈〈cstate … c,current ? (ctape … c)〉, 
340              ntrans M 〈cstate … c,current ? (ctape … c)〉〉) [% //] * #tup #Htup
341 (* tup is in the graph *)
342 cut (mem ? tup (graph_enum ?? (ntrans M)))
343   [@memb_to_mem >Htup @(graph_enum_complete … (ntrans M)) %] #Hingraph
344 (* tupe target = 〈qout,cout,m〉 *)
345 lapply (decomp_target ? (ntrans M 〈cstate … c,current ? (ctape … c)〉))
346 * #qout * #cout * #m #Htg >Htg in Htup; #Htup
347 (* new config *)
348 cut (step FinBool M c = mk_config ?? qout (tape_move ? (tape_write ? (ctape … c) cout) m))
349   [>(config_expand … c) whd in ⊢ (??%?); (* >Htg ?? why not?? *)
350    cut (trans ? M 〈cstate  … c, current ? (ctape … c)〉 = 〈qout,cout,m〉) [<Htg %] #Heq1 
351    >Heq1 %] #Hstep
352 (* new state *)
353 cut (cstate ?? (step FinBool M c) = qout) [>Hstep %] #Hnew_state
354 (* new tape *)
355 cut (ctape ?? (step FinBool M c) = tape_move ? (tape_write ? (ctape … c) cout) m)
356   [>Hstep %] #Hnew_tape
357 lapply(H (bits_of_state ? (nhalt M) (cstate ?? c)) 
358          (low_char (current ? (ctape ?? c)))
359          (tail ? (table_TM ? (graph_enum ?? (ntrans M)) (nhalt M)))
360          ??????)
361 [<Htable
362  lapply(list_to_table … (nhalt M) …Hingraph) * #ll * #lr #Htable1 %{ll} 
363  %{(((bits_of_state ? (nhalt M) qout)@[low_char cout;low_mv m])@lr)} 
364  >Htable1 @eq_f <associative_append @eq_f2 // >Htup
365  whd in ⊢ (??%?); @eq_f >associative_append %
366 |>Ht1 >obj_low_tapes >map_list_of_tape elim (list_of_tape ??) 
367   [#b @False_ind | #b #tl #Hind #a * [#Ha >Ha //| @Hind]]
368 |@sym_eq @Htable
369 |>Ht1 %
370 |%{(bits_of_state ? (nhalt M) (cstate ?? c))} %{(low_char (current ? (ctape ?? c)))}
371  % [% [% [// | cases (current ??) normalize [|#b] % #Hd destruct (Hd)]
372       |>length_map whd in match (length ??); @eq_f //]
373    |//]
374 |>Ht1 >cfg_low_tapes //] -H #H 
375 lapply(H (bits_of_state … (nhalt M) qout) (low_char … cout) 
376          (low_mv … m) tup ? Hingraph)
377   [>Htup whd in ⊢ (??%?); @eq_f >associative_append %] -H
378 #Ht2 @(eq_vec ? 3 ?? (niltape ?) ?) >Ht2 #i #Hi 
379 cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
380   [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
381     [cases (le_to_or_lt_eq … (le_S_S_to_le … Hi)) -Hi #Hi
382       [@False_ind /2/
383       |>Hi >obj_low_tapes >nth_change_vec //
384        >Ht1 >obj_low_tapes >Hstep @map_action 
385       ]
386     |>Hi >cfg_low_tapes >nth_change_vec_neq 
387       [|% whd in ⊢ (??%?→?);  #H destruct (H)]
388      >nth_change_vec // >Hnew_state @eq_f @eq_f >Hnew_tape 
389      @eq_f2 [|2:%] >Ht1 >obj_low_tapes >map_move_mono >low_char_current %
390     ]
391   |(* program tapes do not change *)
392    >Hi >prg_low_tapes 
393    >nth_change_vec_neq [|% whd in ⊢ (??%?→?);  #H destruct (H)]
394    >nth_change_vec_neq [|% whd in ⊢ (??%?→?);  #H destruct (H)]
395    >Ht1 >prg_low_tapes //
396   ]
397 qed.