]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/inject.ma
ecb854b4a13a3900acefdacfe1366c5ce4b6763d
[helm.git] / matita / matita / lib / turing / inject.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 include "turing/turing.ma".
13
14 (******************* inject a mono machine into a multi tape one **********************)
15 definition inject_trans ≝ λsig,states:FinSet.λn,i:nat.
16   λtrans:states × (option sig) → states  × (option sig × move).
17   λp:states × (Vector (option sig) (S n)).
18   let 〈q,a〉 ≝ p in
19   let 〈nq,na〉 ≝ trans 〈q,nth i ? a (None ?)〉 in
20   〈nq, change_vec ? (S n) (null_action ? n) na i〉.
21
22 definition inject_TM ≝ λsig.λM:TM sig.λn,i.
23   mk_mTM sig n
24     (states ? M)
25     (inject_trans sig (states ? M) n i ?) (* (trans sig M))*)
26     (start ? M)
27     (halt ? M).
28 (* ????? *)
29 lapply (trans sig M)  #trans #x lapply (trans x) * *
30 #s #a #m % [ @s | % [ @a | @m ] ]
31 qed.
32
33 lemma current_chars_change_vec: ∀sig,n,v,a,i. i < S n →
34    current_chars sig ? (change_vec ? (S n) v a i) =
35    change_vec ? (S n)(current_chars ?? v) (current ? a) i.
36 #sig #n #v #a #i #Hi @(eq_vec … (None ?)) #i0 #Hi0
37 change with (vec_map ?????) in match (current_chars ???);
38 <(nth_vec_map … (niltape ?))
39 cases (decidable_eq_nat i i0) #Hii0
40 [ >Hii0 >nth_change_vec // >nth_change_vec //
41 | >nth_change_vec_neq // >nth_change_vec_neq // @nth_vec_map ]
42 qed.
43
44 lemma inject_trans_def :∀sig:FinSet.∀n,i:nat.i < S n → 
45   ∀M,v,s,a,sn,an,mn.
46   trans sig M 〈s,a〉 = 〈sn,an,mn〉 → 
47   cic:/matita/turing/turing/trans.fix(0,2,9) sig n (inject_TM sig M n i) 〈s,change_vec ? (S n) v a i〉 = 
48     〈sn,change_vec ? (S n) (null_action ? n) 〈an,mn〉 i〉.
49 #sig #n #i #Hlt #M #v #s #a #sn #an #mn #Htrans
50 whd in ⊢ (??%?); >nth_change_vec // >Htrans //
51 qed.
52
53 lemma inj_ter: ∀A,B,C.∀p:A×B×C. 
54   p = 〈\fst (\fst p), \snd (\fst p), \snd p〉.
55 // qed.
56
57 lemma inject_step : ∀sig,n,M,i,q,t,nq,nt,v. i < S n →
58   step sig M (mk_config ?? q t) = mk_config ?? nq nt → 
59   cic:/matita/turing/turing/step.def(12) sig n (inject_TM sig M n i) 
60     (mk_mconfig ?? n q (change_vec ? (S n) v t i)) 
61   = mk_mconfig ?? n nq (change_vec ? (S n) v nt i).
62 #sig #n #M #i #q #t #nq #nt #v #lein whd in ⊢ (??%?→?);
63 whd in match (step ????); >(current_chars_change_vec … lein)
64 >(inj_ter … (trans sig M ?)) whd in ⊢ (??%?→?); #H
65 >(inject_trans_def sig n i lein M …) 
66 [|>(eq_pair_fst_snd ?? (trans sig M 〈q,current sig t〉))
67   >(eq_pair_fst_snd ?? (\fst (trans sig M 〈q,current sig t〉))) %
68 | *: skip ]
69 whd in ⊢ (??%?); @eq_f2 [destruct (H) // ]
70 @(eq_vec … (niltape ?)) #i0 #lei0n
71 cases (decidable_eq_nat … i i0) #Hii0
72 [ >Hii0 >nth_change_vec // >tape_move_multi_def >pmap_change >nth_change_vec // destruct (H) //
73 | >nth_change_vec_neq // >tape_move_multi_def  >pmap_change >nth_change_vec_neq //
74   <tape_move_multi_def >tape_move_null_action //
75 ]
76 qed. 
77
78 lemma halt_inject: ∀sig,n,M,i,x.
79   cic:/matita/turing/turing/halt.fix(0,2,9) sig n (inject_TM sig M n i) x
80    = halt sig M x.
81 // qed.
82
83 lemma de_option: ∀A,a,b. Some A a = Some A b → a = b. 
84 #A #a #b #H destruct //
85 qed. 
86
87 lemma loop_inject: ∀sig,n,M,i,k,ins,int,outs,outt,vt.i < S n → 
88  loopM sig M k (mk_config ?? ins int) = Some ? (mk_config ?? outs outt) → 
89  cic:/matita/turing/turing/loopM.def(13) sig n (inject_TM sig M n i) k (mk_mconfig ?? n ins (change_vec ?? vt int i))
90   =Some ? (mk_mconfig sig ? n outs (change_vec ?? vt outt i)).
91 #sig #n #M #i #k elim k -k
92   [#ins #int #outs #outt #vt #Hin normalize in ⊢ (%→?); #H destruct
93   |#k #Hind #ins #int #outs #outt #vt #Hin whd in ⊢ (??%?→??%?);
94    >halt_inject whd in match (cstate ????);
95    cases (true_or_false (halt sig M ins)) #Hhalt >Hhalt 
96    whd in ⊢ (??%?→??%?);
97     [#H @eq_f whd in ⊢ (??%?); lapply (de_option ??? H) -H 
98      whd  in ⊢ (??%?→?); #H @eq_f2  
99       [whd in ⊢ (??%?); destruct (H) //
100       |@(eq_vec … (niltape ?)) #j #lejn
101        cases (true_or_false (eqb i j)) #eqij
102         [>(eqb_true_to_eq … eqij) >nth_change_vec //
103          destruct (H) >nth_change_vec //
104         |destruct (H) //
105         ]
106       ]
107     |>(config_expand … (step ???)) #H <(Hind … H) // 
108      >loopM_unfold @eq_f >inject_step //
109     ]
110   ]
111 qed.
112
113 (*
114 lemma cstate_inject: ∀sig,n,M,i,x. *)
115
116 definition inject_R ≝ λsig.λR:relation (tape sig).λn,i:nat.
117   λv1,v2: (Vector (tape sig) (S n)).
118   R (nth i ? v1 (niltape ?)) (nth i ? v2 (niltape ?)) ∧
119   ∀j. i ≠ j → nth j ? v1 (niltape ?) = nth j ? v2 (niltape ?). 
120
121 (*
122 lemma nth_make : ∀A,i,n,j,a,d. i < n → nth i ? (make_veci A a n j) d = a (j+i).
123 #A #i elim i
124   [#n #j #a #d #ltOn @(lt_O_n_elim … ltOn) <plus_n_O //
125   |#m #Hind #n #j #a #d #Hlt lapply Hlt @(lt_O_n_elim … (ltn_to_ltO … Hlt)) 
126    #p <plus_n_Sm #ltmp @Hind @le_S_S_to_le //  
127   ]
128 qed. *)
129
130 (*
131 lemma mk_config_eq_s: ∀S,sig,s1,s2,t1,t2. 
132   mk_config S sig s1 t1 = mk_config S sig s2 t2 → s1=s2.
133 #S #sig #s1 #s2 #t1 #t2 #H destruct //
134 qed.
135
136 lemma mk_config_eq_t: ∀S,sig,s1,s2,t1,t2. 
137   mk_config S sig s1 t1 = mk_config S sig s2 t2 → s1=s2.
138 #S #sig #s1 #s2 #t1 #t2 #H destruct //
139 qed.
140 *)
141
142 theorem sem_inject: ∀sig.∀M:TM sig.∀R.∀n,i.
143  i≤n → M ⊨ R → inject_TM sig M n i ⊨ inject_R sig R n i. 
144 #sig #M #R #n #i #lein #HR #vt cases (HR (nth i ? vt (niltape ?)))
145 #k * * #outs #outt * #Hloop #HRout @(ex_intro ?? k) 
146 @(ex_intro ?? (mk_mconfig ?? n outs (change_vec ? (S n) vt outt i))) %
147   [whd in ⊢ (??(?????%)?); <(change_vec_same ?? vt i (niltape ?)) in ⊢ (??%?);
148    @loop_inject /2 by refl, trans_eq, le_S_S/
149   |%[>nth_change_vec // @le_S_S //
150     |#j #i >nth_change_vec_neq //
151     ]
152   ]
153 qed.
154
155 theorem acc_sem_inject: ∀sig.∀M:TM sig.∀Rtrue,Rfalse,acc.∀n,i.
156  i≤n → M ⊨ [ acc : Rtrue, Rfalse ] → 
157  inject_TM sig M n i ⊨ [ acc : inject_R sig Rtrue n i, inject_R sig Rfalse n i ]. 
158 #sig #M #Rtrue #Rfalse #acc #n #i #lein #HR #vt cases (HR (nth i ? vt (niltape ?)))
159 #k * * #outs #outt * * #Hloop #HRtrue #HRfalse @(ex_intro ?? k) 
160 @(ex_intro ?? (mk_mconfig ?? n outs (change_vec ? (S n) vt outt i))) % [ %
161   [whd in ⊢ (??(?????%)?); <(change_vec_same ?? vt i (niltape ?)) in ⊢ (??%?);
162    @loop_inject /2 by refl, trans_eq, le_S_S/
163   |#Hqtrue %
164     [>nth_change_vec /2 by le_S_S/
165     |#j #Hneq >nth_change_vec_neq //
166     ] ]
167   |#Hqfalse %
168     [>nth_change_vec /2 by le_S_S/ @HRfalse @Hqfalse
169     |#j #Hneq >nth_change_vec_neq //
170     ] ]
171 qed.