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