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