]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/normalTM.ma
semantics of unistep
[helm.git] / matita / matita / lib / turing / multi_universal / normalTM.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/multi_universal/alphabet.ma".
13 include "turing/mono.ma".
14
15 (************************* turning DeqMove into a DeqSet **********************)
16 definition move_eq ≝ λm1,m2:move.
17   match m1 with
18   [R ⇒ match m2 with [R ⇒ true | _ ⇒ false]
19   |L ⇒ match m2 with [L ⇒ true | _ ⇒ false]
20   |N ⇒ match m2 with [N ⇒ true | _ ⇒ false]].
21
22 lemma move_eq_true:∀m1,m2.
23   move_eq m1 m2 = true ↔ m1 = m2.
24 *
25   [* normalize [% #_ % |2,3: % #H destruct ]
26   |* normalize [1,3: % #H destruct |% #_ % ]
27   |* normalize [1,2: % #H destruct |% #_ % ]
28 qed.
29
30 definition DeqMove ≝ mk_DeqSet move move_eq move_eq_true.
31
32 unification hint 0 ≔ ;
33     X ≟ DeqMove
34 (* ---------------------------------------- *) ⊢ 
35     move ≡ carr X.
36
37 unification hint  0 ≔ m1,m2; 
38     X ≟ DeqMove
39 (* ---------------------------------------- *) ⊢ 
40     move_eq m1 m2 ≡ eqb X m1 m2.
41
42
43 (************************ turning DeqMove into a FinSet ***********************)
44 definition move_enum ≝ [L;R;N].
45
46 lemma move_enum_unique: uniqueb ? [L;R;N] = true.
47 // qed.
48
49 lemma move_enum_complete: ∀x:move. memb ? x [L;R;N] = true.
50 * // qed.
51
52 definition FinMove ≝ 
53   mk_FinSet DeqMove [L;R;N] move_enum_unique move_enum_complete.
54
55 unification hint  0 ≔ ; 
56     X ≟ FinMove
57 (* ---------------------------------------- *) ⊢ 
58     move ≡ FinSetcarr X.
59
60 (*************************** normal Turing Machines ***************************)
61
62 (* A normal turing machine is just an ordinary machine where:
63   1. the tape alphabet is bool
64   2. the finite state are supposed to be an initial segment of the natural 
65      numbers. 
66   Formally, it is specified by a record with the number n of states, a proof
67   that n is positive, the transition function and the halting function.
68 *)
69
70 definition trans_source ≝ λn.
71   FinProd (initN n) (FinOption FinBool).
72   
73 definition trans_target ≝ λn.
74   FinProd (FinProd (initN n) (FinOption FinBool)) FinMove.
75
76 record normalTM : Type[0] ≝ 
77 { no_states : nat;
78   pos_no_states : (0 < no_states); 
79   ntrans : trans_source no_states → trans_target no_states;
80   nhalt : initN no_states → bool
81 }.
82
83 lemma decomp_target : ∀n.∀tg: trans_target n. 
84   ∃qout,cout,m. tg = 〈qout,cout,m〉.
85 #n * * #q #c #m %{q} %{c} %{m} //
86 qed.
87
88 (* A normal machine is just a special case of Turing Machine. *)
89
90 definition nstart_state ≝ λM.mk_Sig ?? 0 (pos_no_states M).
91
92 definition normalTM_to_TM ≝ λM:normalTM.
93   mk_TM FinBool (initN (no_states M)) 
94    (ntrans M) (nstart_state M) (nhalt M).
95
96 coercion normalTM_to_TM.
97
98 definition nconfig ≝ λn. config FinBool (initN n).
99
100 (* A normal machine has a non empty graph *)
101
102 definition sample_input: ∀M.trans_source (no_states M).
103 #M % [@(nstart_state M) | %]
104 qed.
105
106 lemma nTM_nog: ∀M. graph_enum ?? (ntrans M) ≠ [ ].
107 #M % #H lapply(graph_enum_complete ?? (ntrans M) (sample_input M) ? (refl ??))
108 >H normalize #Hd destruct
109 qed.
110
111 (******************************** tuples **************************************)
112
113 (* By general results on FinSets we know that every function f between two 
114 finite sets A and B can be described by means of a finite graph of pairs
115 〈a,f a〉. Hence, the transition function of a normal turing machine can be
116 described by a finite set of tuples 〈i,c〉,〈j,action〉〉 of the following type:
117   (Nat_to n × (option FinBool)) × (Nat_to n × (option FinBool) × move).  
118 Unfortunately this description is not suitable for a Universal Machine, since
119 such a machine must work with a fixed alphabet, while the size on n is unknown.
120 Hence, we must pass from natural numbers to a representation for them on a 
121 finitary, e.g. binary, alphabet. In general, we shall associate
122 to a pair 〈〈i,x〉,〈j,y,m〉〉 a tuple with the following syntactical structure
123            | w_i x w_j y m
124 where 
125 1. "|" is a special character used to separate tuples
126 2. w_i and w_j are list of booleans representing the states $i$ and $j$; 
127 3. x and y are encoding 
128 4. finally, m ...
129 *)
130
131 definition mk_tuple ≝ λqin,cin,qout,cout,mv.
132   bar::qin@cin::qout@[cout;mv].
133
134 definition tuple_TM : nat → list unialpha → Prop ≝ 
135  λn,t.∃qin,cin,qout,cout,mv.
136  only_bits qin ∧ only_bits qout ∧ cin ≠ bar ∧ cout ≠ bar ∧ mv ≠ bar ∧
137  |qin| = n ∧ |qout| = n ∧
138  t = mk_tuple qin cin qout cout mv. 
139
140 (***************************** state encoding *********************************)
141 (* p < n is represented with a list of bits of lenght n with the p-th bit from 
142 left set to 1. An additional intial bit is set to 1 if the state is final and
143 to 0 otherwise. *)
144  
145 let rec to_bitlist n p: list bool ≝
146   match n with [ O ⇒ [ ] | S q ⇒ (eqb p q)::to_bitlist q p].
147   
148 let rec from_bitlist l ≝
149   match l with 
150   [ nil ⇒ 0 (* assert false *)
151   | cons b tl ⇒ if b then |tl| else from_bitlist tl].
152
153 lemma bitlist_length: ∀n,p.|to_bitlist n p| = n.
154 #n elim n normalize // 
155 qed.
156
157 lemma bitlist_inv1: ∀n,p.p<n → from_bitlist (to_bitlist n p) = p.
158 #n elim n normalize -n
159   [#p #abs @False_ind /2/
160   |#n #Hind #p #lepn 
161    cases (le_to_or_lt_eq … (le_S_S_to_le … lepn))
162     [#ltpn lapply (lt_to_not_eq … ltpn) #Hpn
163      >(not_eq_to_eqb_false … Hpn) normalize @Hind @ltpn
164     |#Heq >(eq_to_eqb_true … Heq) normalize <Heq //
165     ]
166   ]
167 qed.
168
169 lemma bitlist_lt: ∀l. 0 < |l| → from_bitlist l < |l|.
170 #l elim l normalize // #b #tl #Hind cases b normalize //
171 #Htl cases (le_to_or_lt_eq … (le_S_S_to_le … Htl)) -Htl #Htl
172   [@le_S_S @lt_to_le @Hind //  
173   |cut (tl=[ ]) [/2 by append_l2_injective/] #eqtl >eqtl @le_n
174   ]
175 qed.
176
177 definition nat_of: ∀n. Nat_to n → nat.
178 #n normalize * #p #_ @p
179 qed. 
180
181 definition bits_of_state ≝ λn.λh:Nat_to n → bool.λs:Nat_to n.
182    map ? unialpha (λx.bit x) (h s::(to_bitlist n (nat_of n s))).
183
184 lemma only_bits_bits_of_state : ∀n,h,p. only_bits (bits_of_state n h p).
185 #n #h #p #x whd in match (bits_of_state n h p);
186 #H cases H -H 
187   [#H >H %
188   |elim (to_bitlist n (nat_of n p))
189     [@False_ind |#b #l #Hind #H cases H -H #H [>H % |@Hind @H ]]
190   ]
191 qed.
192
193 (******************************** action encoding *****************************)
194
195 definition low_char ≝ λc. 
196   match c with 
197     [ None ⇒ null
198     | Some b ⇒ bit b
199     ].
200     
201 definition low_mv ≝ λm. 
202   match m with 
203   [ R ⇒ bit true
204   | L ⇒ bit false
205   | N ⇒ null
206   ].
207
208 (******************************** tuple encoding ******************************)
209 definition tuple_type ≝ λn.
210  (Nat_to n × (option FinBool)) × (Nat_to n × (option FinBool) × move).  
211
212 definition tuple_encoding ≝ λn.λh:Nat_to n→bool. 
213   λp:tuple_type n.
214   let 〈inp,outp〉 ≝ p in
215   let 〈q,a〉 ≝ inp in
216   let 〈qn,an,m〉 ≝ outp in
217   let qin ≝ bits_of_state n h q in
218   let qout ≝ bits_of_state n h qn in
219   let cin ≝ low_char a in
220   let cout ≝ low_char an in
221   let mv ≝ low_mv m in
222   mk_tuple qin cin qout cout mv.
223
224 lemma is_tuple: ∀n,h,p. tuple_TM (S n) (tuple_encoding n h p).
225 #n #h * * #q #a * * #qn #an #m
226 %{(bits_of_state n h q)} %{(low_char a)} 
227 %{(bits_of_state n h qn)} %{(low_char an)} %{(low_mv m)} 
228 % // % 
229   [%[%[%[%[% /2/ |% cases a normalize [|#b] #H destruct]
230         |% cases an normalize [|#b] #H destruct]
231       |% cases m normalize #H destruct]
232     |>length_map normalize @eq_f //]
233   |>length_map normalize @eq_f //]
234 qed. 
235
236 definition tuple_length ≝ λn.2*n+4.
237
238 lemma length_of_tuple: ∀n,t. tuple_TM n t → 
239   |t| = tuple_length n.
240 #n #t * #qin * #cin * #qout * #cout * #mv *** #_ #Hqin #Hqout #eqt 
241 >eqt normalize >length_append >Hqin -Hqin normalize >length_append 
242 normalize >Hqout //
243 qed.
244
245 definition tuples_list ≝ λn.λh.map … (λp.tuple_encoding n h p).
246
247
248
249