]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_to_mono/step.ma
The step machine (draft)
[helm.git] / matita / matita / lib / turing / multi_to_mono / step.ma
1
2 include "turing/multi_to_mono/exec_moves.ma".
3
4 definition Mono_realize ≝ Realize.
5
6 definition to_sig ≝ λQ,sig.λc:sig_ext (TA Q sig).
7   match c with 
8   [None ⇒ None ?
9   |Some a ⇒ match a with 
10     [inl x ⇒ None ? (* this is not possible *)
11     |inr x ⇒ Some ? x]].
12
13 definition to_sig_inv ≝ λQ,sig.λc1:(option sig) × move.λc2.
14   match (fst ?? c1) with 
15   [None ⇒ c2
16   |Some a ⇒ Some ? (inr Q ? a)].
17
18 definition transf ≝ λQ,sig:FinSet.λn.
19  λt: Q × (Vector (option sig) n) → Q × (Vector ((option sig) × move) n).
20  λa:MTA (HC Q n) sig (S n).
21   let qM ≝ nth n ? (vec … a) (blank ?) in
22   let a1 ≝ resize_vec ? (S n) a n (blank ?) in
23   let a2 ≝ vec_map ?? (to_sig ? sig) n a1 in
24   match qM with 
25   [None ⇒ all_blanks (TA (HC Q n) sig) (S n)
26   |Some p ⇒ 
27     match p with 
28     [inl p1 ⇒ 
29       let 〈q1,actions〉 ≝ t 〈fst ?? p1, a2〉 in
30       let moves ≝ vec_map ?? (snd ??) n actions in 
31       let new_chars ≝ pmap_vec ??? (to_sig_inv (HC Q n) ?) n actions a1 in
32       vec_cons_right ? (Some ? (inl ?? 〈q1,moves〉)) n new_chars
33     |inr p2 ⇒ all_blanks (TA (HC Q n) sig) (S n)]
34   ].
35
36 lemma transf_eq: ∀Q,sig,n,t.∀a:MTA (HC Q n) sig (S n).
37   ∀a1,a2,q,m,q1,actions,moves,new_chars.
38   nth n ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉) →
39   a1 = resize_vec ? (S n) a n (blank ?) →
40   a2 = vec_map ?? (to_sig ? sig) n a1 →
41   t 〈q, a2〉 = 〈q1,actions〉 →
42   moves = vec_map ?? (snd ??) n actions → 
43   new_chars = pmap_vec ??? (to_sig_inv (HC Q n) ?) n actions a1  →
44   transf Q sig n t a = 
45     vec_cons_right ? (Some ? (inl ?? 〈q1,moves〉)) n new_chars.
46 #Q #sig #n #t #a #a1 #a2 #q #m #q1 #actions #moves #new_chars
47 #Hnth #Ha1 #Ha2 #Ht #Hmoves #Hnew_chars
48 whd in ⊢ (??%?); >Hnth whd in ⊢ (??%?); <Ha1 <Ha2 >Ht
49 @eq_cons_right [<Hmoves // |<Hnew_chars // ]
50 qed.
51   
52 lemma transf_eq_ex: ∀Q,sig,n,t.∀a:MTA (HC Q n) sig (S n).
53   ∀q,m.nth n ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉)→
54   ∃a1.a1 = resize_vec ? (S n) a n (blank ?) ∧
55   ∃a2. a2 = vec_map ?? (to_sig ? sig) n a1 ∧
56   ∃q1,actions. t 〈q, a2〉 = 〈q1,actions〉 ∧
57   ∃moves.moves = vec_map ?? (snd ??) n actions ∧ 
58   ∃new_chars.new_chars = pmap_vec ??? (to_sig_inv (HC Q n) ?) n actions a1 ∧
59   transf Q sig n t a = 
60     vec_cons_right ? (Some ? (inl ?? 〈q1,moves〉)) n new_chars.
61 #Q #sig #n #t #a #q #m #H 
62 %{(resize_vec ? (S n) a n (blank ?))} % [%] 
63 % [2:% [%] | skip] % [2: % [ 2: % [@eq_pair_fst_snd ] |skip] | skip]
64 % [2:% [%] | skip] % [2:% [%] | skip] whd in ⊢ (??%?); >H whd in ⊢ (??%?);
65 >(eq_pair_fst_snd … (t …)) % 
66 qed.
67
68 (*
69 lemma nth_transf: ∀Q,sig,n,t.∀a:MTA (HC Q n) sig (S n).
70   ∀a1,a2,q,m,q1,actions,moves.
71   nth n ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉) →
72   a1 = resize_vec ? (S n) a n (blank ?) →
73   a2 = vec_map ?? (to_sig ? sig) n a1 →
74   t 〈q, a2〉 = 〈q1,actions〉 →
75   moves = vec_map ?? (snd ??) n actions → 
76   nth n ? (vec … (transf Q sig n t a)) (blank ?) = Some ? (inl ?? 〈q1,moves〉).
77 #Q #sig #n #t #a #a1 #a2 #q #m #q1 #actions #moves
78 #Hnth #Ha1 #Ha2 #Ht #Hmoves
79 *)
80
81 (*********************************** stepM ************************************)
82 definition optf ≝ λA,B:Type[0].λf:A →B.λd,oa.
83   match oa with 
84   [ None ⇒ d
85   | Some a ⇒ f a ].
86
87 lemma optf_Some : ∀A,B,f,a,d. optf A B f d (Some A a) = f a.
88 // qed.
89
90 definition stepM ≝ λQ,sig,n,trans.
91   writef ? (optf ?? (transf Q sig (S n) trans) (all_blanks …)) ·
92   exec_moves Q sig (S n) (S n).
93
94 let rec to_sig_map Q sig l on l ≝
95   match l with 
96   [ nil ⇒ nil ?
97   | cons a tl ⇒ match to_sig Q sig a with 
98     [ None ⇒ nil ?
99     | Some c ⇒ c::(to_sig_map Q sig tl)]].
100
101
102 definition to_sig_tape ≝ λQ,sig,t.
103   match t with
104   [ niltape  ⇒ niltape ?
105   | leftof a l ⇒ match to_sig Q sig a with 
106     [ None ⇒ niltape ?
107     | Some x ⇒ leftof ? x (to_sig_map Q sig l) ]
108   | rightof a l ⇒ match to_sig Q sig a with 
109     [ None ⇒ niltape ?
110     | Some x ⇒ rightof ? x (to_sig_map Q sig l) ]
111   | midtape ls a rs ⇒ match to_sig Q sig a with 
112     [ None ⇒ niltape ?
113     | Some x ⇒ midtape ? (to_sig_map Q sig ls) x (to_sig_map Q sig rs) ]
114   ].
115
116 definition rb_tapes ≝ λQ,sig,n,ls.λa:MTA Q sig (S n).λrs.
117   vec_map ?? (to_sig_tape ??) n (readback ? (S n) ls (vec … a) rs n).
118
119 (* q0 is a default value *)
120 definition get_state ≝ λQ,sig,n.λa:MTA (HC Q n) sig (S n).λq0.
121   match nth n ? (vec … a) (blank ?) with 
122   [ None ⇒ q0 (* impossible *)
123   | Some qM ⇒ match qM with
124     [inl qM1 ⇒ fst ?? qM1 
125     |inr _ ⇒ q0 (* impossible *) ]].
126     
127 definition get_chars ≝ λQ,sig,n.λa:MTA (HC Q n) sig (S n).
128   let a1 ≝ resize_vec ? (S n) a n (blank ?) in
129   vec_map ?? (to_sig ? sig) n a1.
130   
131 lemma get_chars_def : ∀Q,sig,n.∀a:MTA (HC Q n) sig (S n).
132   get_chars … a =
133     vec_map ?? (to_sig ? sig) n (resize_vec ? (S n) a n (blank ?)).
134 // qed. 
135
136 include "turing/turing.ma".
137
138 definition readback_config ≝
139   λQ,sig,n,q0,ls.λa:MTA (HC Q (S n)) sig (S (S n)).λrs.
140   mk_mconfig sig Q n 
141     (get_state Q sig (S n) a q0)
142     (rb_tapes (HC Q (S n)) sig (S n) ls a rs).
143
144 lemma state_readback : ∀Q,sig,n,q0,ls,a,rs.
145   cstate … (readback_config Q sig n q0 ls a rs) =
146     get_state Q sig (S n) a q0.
147 // qed.
148
149 lemma tapes_readback : ∀Q,sig,n,q0,ls,a,rs.
150   ctapes … (readback_config Q sig n q0 ls a rs) =
151     rb_tapes (HC Q (S n)) sig (S n) ls a rs.
152 // qed.
153     
154 definition R_stepM ≝ λsig.λn.λM:mTM sig n.λt1,t2.
155 ∀a,ls,rs. 
156   t1 = midtape ? ls a rs → 
157   (∀i.regular_trace (TA (HC (states … M) (S n)) sig) (S(S n)) a ls rs i) → 
158   is_head ?? (nth (S n) ? (vec … a) (blank ?)) = true →  
159   no_head_in … ls →
160   no_head_in … rs →
161   ∃ls1,a1,rs1. 
162    t2 = midtape (MTA (HC (states … M) (S n)) sig (S(S n))) ls1 a1 rs1 ∧
163    (∀i.regular_trace (TA (HC (states … M) (S n)) sig) (S(S n)) a1 ls1 rs1 i) ∧
164    readback_config ? sig n (start … M) ls1 a1 rs1 =
165      step sig n M (readback_config ? sig n (start … M) ls a rs).
166
167 lemma nth_vec_map_lt : 
168   ∀A,B,f,i,n.∀v:Vector A n.∀d1,d2.i < n →
169   nth i ? (vec_map A B f n v) d1 = f (nth i ? v d2).
170 #A #B #f #i #n #v #d1 #d2 #ltin >(nth_default B i n ? d1 (f d2) ltin) @sym_eq @nth_vec_map 
171 qed.
172
173 lemma ctapes_mk_config : ∀sig,Q,n,s,t.
174   ctapes sig Q n (mk_mconfig sig Q n s t) = t.
175 // qed.
176
177 lemma cstate_rb: ∀sig,n.∀M:mTM sig n.∀ls,a,rs.∀q,m.
178   nth (S n) ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉) →
179   cstate sig (states sig n M) n
180       (readback_config (states sig n M) sig n (start sig n M) ls a rs) = q.
181 #sig #n #M #ls #a #rs #q #m #H
182 >state_readback whd in ⊢ (??%?); >H %
183 qed.
184
185 axiom eq_cstate_get_state: ∀sig,n.∀M:mTM sig n.∀ls,a,rs.
186   is_head ?? (nth (S n) ? (vec … a) (blank ?)) = true →
187   cstate sig (states sig n M) n
188       (readback_config (states sig n M) sig n (start sig n M) ls a rs) = 
189   get_state (states sig n M) sig (S n) a (start sig n M). 
190
191 axiom eq_current_chars_resize: ∀sig,n.∀M:mTM sig n.∀ls,a,rs.
192   current_chars sig n
193    (ctapes sig (states sig n M) n
194     (readback_config (states sig n M) sig n (start sig n M) ls a rs)) =
195   get_chars … a.
196   
197 (*
198 lemma rb_trans : ∀sig,Q,n.∀M:mTM sig n.∀ls,a,rs,q,m.
199   nth (S n) ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉) →
200   transf (states sig n M) sig (S n) (trans sig n M) a =
201   trans sig n M 〈q,get_chars … a〉.
202 *)
203
204 lemma nth_pmap_lt :
205   ∀A,B,C,f,i,n.∀v1:Vector A n.∀v2:Vector B n.∀d1,d2,d3.i < n→
206   f (nth i ? v1 d1) (nth i ? v2 d2) = nth i ? (pmap_vec A B C f n v1 v2) d3.
207 #A #B #C #f #i elim i
208 [ *
209   [ #v1 #v2 #d1 #d2 #d3 #Hlt @False_ind /2/  
210   | #n0 #v1 #v2 #d1 #d2 #d3 #_ >(vec_expand … v1)
211     >(vec_expand … v2) >(nth_default … d3 (f d1 d2)) [% | // ]]
212 | #i0 #IH *
213   [ #v1 #v2 #d1 #d2 #d3 #Hlt @False_ind /2/  
214   | #n #v1 #v2 #d1 #d2 #d3 #Hlt>(vec_expand … v1) 
215      >(vec_expand … v2) 
216      whd in match (nth …d1); whd in match (tail ??); 
217      whd in match (nth … d2); whd in match (tail B ?);
218      whd in match (nth … d3); whd in match (tail C ?); 
219      <(IH n ?? d1 d2 d3) [2:@le_S_S_to_le @Hlt] %
220   ] ]
221 qed.
222
223 lemma tape_move_mono_def : ∀sig,t,a,m.
224   tape_move_mono sig t 〈a,m〉 = tape_move sig (tape_write sig t a) m.
225 // qed.
226
227 axiom to_sig_move : ∀Q,sig,n,t,m.
228   to_sig_tape (HC Q (S n)) sig
229   (tape_move (sig_ext (TA (HC Q (S n)) sig)) t m)
230   = tape_move sig (to_sig_tape ?? t) m.
231   
232 definition to_sig_conv :∀Q,sig:FinSet.∀n.
233  option sig → option (sig_ext (TA (HC Q (S n)) sig))
234 ≝ λQ,sig,n,c.
235   match c with 
236   [None ⇒ None ?
237   |Some a ⇒ Some ? (Some ? (inr (HC Q (S n)) ? a))].
238  
239 axiom to_sig_write : ∀Q,sig,n,t,c.
240   to_sig_tape (HC Q (S n)) sig
241   (tape_write ? t (to_sig_conv ??? c))
242    = tape_write sig (to_sig_tape ?? t) c.
243
244 axiom rb_write: ∀sig,n,ls,a,rs,i,c1,c2.
245   rb_trace_i ? n ls c1 rs i =
246   tape_write ? (rb_trace_i sig n ls a rs i) c2.
247    
248 lemma sem_stepM : ∀sig,n.∀M:mTM sig n.
249   stepM (states sig n M) sig n (trans sig n M) ⊨
250     R_stepM sig n M.
251 #sig #n #M 
252 @(sem_seq_app … (sem_writef … ) (sem_exec_moves … (le_n ?)))
253 #tin #tout * #t1 * whd in ⊢ (%→?); #Hwrite #Hmoves 
254 #a #ls #rs #Htin #H1 #H2 #H3 #H4 >Htin in Hwrite; #Hwrite
255 lapply (Hwrite … (refl …)) -Hwrite whd in match (right ??); whd in match (left ??);
256 >optf_Some #Ht1
257 cut (∃q,m. nth (S n) ? (vec … a) (blank ?) = Some ? (inl ?? 〈q,m〉))
258     [lapply H2 cases (nth (S n) ? (vec … a) (blank ?)) 
259       [whd in ⊢ (??%?→?); #H destruct (H)
260       |* #x whd in ⊢ (??%?→?); #H destruct (H) cases x #q #m %{q} %{m} %
261       ]
262     ] * #q * #m #HaSn
263 (* 
264 lapply (transf_eq … HaSn (refl ??) (refl ??) (eq_pair_fst_snd …) (refl ??) (refl ??))
265 *)
266 lapply (Hmoves … Ht1 ?? H3 H4)
267   [>(transf_eq … HaSn (refl ??) (refl ??) (eq_pair_fst_snd …) (refl ??) (refl ??))
268    >nth_cons_right %
269   | (* regularity is preserved *) @daemon
270   |* #ls1 * #a1 * #rs1 * * * #Htout #Hreg #Hrb #HtrSn
271    lapply (HtrSn (S n) (le_n ?) (le_n ?)) -HtrSn #HtrSn
272    cut (nth (S n) ? (vec … a1) (blank ?) = 
273         nth (S n ) ? (vec … (transf (states sig n M) sig (S n) (trans sig n M) a)) (blank ?))
274     [@daemon (* from HtrSn *)] #Ha1
275    %{ls1} %{a1} %{rs1} 
276    %[%[@Htout |@Hreg]
277     |(* commutation *)
278      lapply(transf_eq_ex …  (trans sig n M) … HaSn)
279      * #c1 * #Hc1 * #c2 * #Hc2 * #q1 * #actions *
280      #Htrans * #moves * #Hmoves * #new_chars * #Hnew_chars #Htransf
281      @mconfig_eq 
282       [(* state *) >state_readback whd in match (step ????);
283        >(cstate_rb … HaSn) >eq_current_chars_resize >get_chars_def
284        <Hc1 <Hc2 >Htrans whd in ⊢ (???%);
285        whd in ⊢ (??%?); >Ha1 >HaSn >Htransf >nth_cons_right %
286       |>tapes_readback whd in match (step ????);
287        >(cstate_rb … HaSn) >eq_current_chars_resize >get_chars_def
288        <Hc1 <Hc2 >Htrans >ctapes_mk_config 
289        @(eq_vec … (niltape ?)) #i #lti
290        >nth_vec_map_lt [2:@lti |3:@niltape]
291        >Hrb <nth_pmap_lt [2:@lti|3:@N|4:@niltape]
292        >tapes_readback
293        >Htransf whd in match (vec_moves ?????);
294        >get_moves_cons_right >resize_id [2:@(len ?? moves)]
295        (* lhs *)
296        <nth_pmap_lt [2:@lti|3:%[@None|@N]|4:@niltape]
297        >nth_vec_map_lt [2:@lti |3:@niltape]
298        >(eq_pair_fst_snd … (nth i ? actions ?))
299        >tape_move_mono_def
300        cut (snd ?? (nth i ? actions 〈None sig,N〉) = nth i ? moves N)
301          [>Hmoves @nth_vec_map] #Hmoves1 >Hmoves1
302        >(nth_readback … lti) >(nth_readback … lti)
303        >to_sig_move @eq_f2 [2://]
304        <to_sig_write @eq_f @rb_write (* finto *)
305       ]
306     ]
307   ]
308 qed.
309        
310        
311   
312   
313   
314   
315
316
317
318
319
320
321
322
323
324