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