]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/universal/trans_to_tuples.ma
e40a0bbc8d5d5b446f7cdfb356154069a690ce9e
[helm.git] / matita / matita / lib / turing / universal / trans_to_tuples.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
14 include "turing/universal/tuples.ma".
15 include "basics/fun_graph.ma".
16
17 (* p < n is represented with a list of bits of lenght n with the
18  p-th bit from left set to 1 *)
19  
20 let rec to_bitlist n p: list bool ≝
21   match n with
22   [ O ⇒ [ ]
23   | S q ⇒ (eqb p q)::to_bitlist q p].
24   
25 let rec from_bitlist l ≝
26   match l with
27   [ nil ⇒ 0 (* assert false *)
28   | cons b tl ⇒ if b then |tl| else from_bitlist tl].
29
30 lemma bitlist_length: ∀n,p.|to_bitlist n p| = n.
31 #n elim n normalize // 
32 qed.
33   
34 lemma bitlist_inv1: ∀n,p.p<n → from_bitlist (to_bitlist n p) = p.
35 #n elim n normalize -n
36   [#p #abs @False_ind /2/
37   |#n #Hind #p #lepn 
38    cases (le_to_or_lt_eq … (le_S_S_to_le … lepn))
39     [#ltpn lapply (lt_to_not_eq … ltpn) #Hpn
40      >(not_eq_to_eqb_false … Hpn) normalize @Hind @ltpn
41     |#Heq >(eq_to_eqb_true … Heq) normalize <Heq //
42     ]
43   ]
44 qed.
45
46 lemma bitlist_lt: ∀l. 0 < |l| → from_bitlist l < |l|.
47 #l elim l normalize // #b #tl #Hind cases b normalize //
48 #Htl cases (le_to_or_lt_eq … (le_S_S_to_le … Htl)) -Htl #Htl
49   [@le_S_S @lt_to_le @Hind //  
50   |cut (tl=[ ]) [/2 by append_l2_injective/] #eqtl >eqtl @le_n
51   ]
52 qed.
53
54 definition nat_of: ∀n. Nat_to n → nat.
55 #n normalize * #p #_ @p
56 qed. 
57
58 definition bits_of_state ≝ λn.λh:Nat_to n → bool.λs:Nat_to n.
59   h s::(to_bitlist n (nat_of n s)).
60
61 definition m_bits_of_state ≝ λn.λh.λp.
62   map ? (unialpha×bool) (λx.〈bit x,false〉) (bits_of_state n h p).
63   
64 lemma no_marks_bits_of_state : ∀n,h,p. no_marks (m_bits_of_state n h p).
65 #n #h #p #x whd in match (m_bits_of_state n h p);
66 #H cases (orb_true_l … H) -H 
67   [#H >(\P H) %
68   |elim (to_bitlist n (nat_of n p))
69     [whd in ⊢ ((??%?)→?); #H destruct 
70     |#b #l #Hind #H cases (orb_true_l … H) -H #H
71       [>(\P H) %
72       |@Hind @H
73       ]
74     ]
75   ]
76 qed.
77
78 lemma only_bits_bits_of_state : ∀n,h,p. only_bits (m_bits_of_state n h p).
79 #n #h #p #x whd in match (m_bits_of_state n h p);
80 #H cases (orb_true_l … H) -H 
81   [#H >(\P H) %
82   |elim (to_bitlist n (nat_of n p))
83     [whd in ⊢ ((??%?)→?); #H destruct 
84     |#b #l #Hind #H cases (orb_true_l … H) -H #H
85       [>(\P H) %
86       |@Hind @H
87       ]
88     ]
89   ]
90 qed.
91
92 definition tuple_type ≝ λn.
93  (Nat_to n × (option FinBool)) × (Nat_to n × (option (FinBool × move))).  
94
95 definition tuple_of_pair ≝ λn.λh:Nat_to n→bool. 
96   λp:tuple_type n.
97   let 〈inp,outp〉 ≝ p in
98   let 〈q,a〉 ≝ inp in
99   let cin ≝ match a with [ None ⇒ null | Some b ⇒ bit b ] in
100   let 〈qn,action〉 ≝ outp in
101   let 〈cout,mv〉 ≝ 
102     match action with 
103     [ None ⇒ 〈null,null〉
104     | Some act ⇒ let 〈na,m〉 ≝ act in 
105       match m with 
106       [ R ⇒ 〈bit na,bit true〉
107       | L ⇒ 〈bit na,bit false〉
108       | N ⇒ 〈bit na,null〉]
109     ] in
110   let qin ≝ m_bits_of_state n h q in
111   let qout ≝ m_bits_of_state n h qn in
112   mk_tuple qin 〈cin,false〉 qout 〈cout,false〉 〈mv,false〉.
113
114 definition WFTuple_conditions ≝ 
115  λn,qin,cin,qout,cout,mv.
116  no_marks qin ∧ no_marks qout ∧ (* queste fuori ? *)
117  only_bits qin ∧ only_bits qout ∧ 
118  bit_or_null cin = true ∧ bit_or_null cout = true ∧ bit_or_null mv = true ∧
119  (cout = null → mv = null) ∧
120  |qin| = n ∧ |qout| = n.
121
122 lemma is_tuple: ∀n,h,p. tuple_TM (S n) (tuple_of_pair n h p).
123 #n #h * * #q #a * #qn #action
124 @(ex_intro … (m_bits_of_state n h q))
125 letin cin ≝ match a with [ None ⇒ null | Some b ⇒ bit b ]
126 @(ex_intro … cin)
127 @(ex_intro … (m_bits_of_state n h qn))
128 letin cout ≝ 
129   match action with 
130   [ None ⇒ null | Some act ⇒ bit (\fst act)]
131 @(ex_intro … cout)
132 letin mv ≝ match action with 
133   [ None ⇒ null
134   | Some act ⇒ 
135       match \snd act with 
136       [ R ⇒ bit true | L ⇒ bit false | N ⇒ null]
137   ]
138 @(ex_intro … mv)
139 %[%[%[%[%[%[%[% /3/ 
140              |whd in match cin ; cases a //
141              ]
142            |whd in match cout; cases action //
143            ]
144          |whd in match mv; cases action //
145           * #b #m cases m //
146          ]
147        |whd in match cout; whd in match mv; cases action
148          [// | #act whd in ⊢ ((??%?)→?); #Hfalse destruct ]
149        ]
150      |>length_map normalize @eq_f //
151      ]
152    |>length_map normalize @eq_f //
153    ]
154  |normalize cases a cases action normalize //
155    [* #c #m cases m %
156    |* #c #m #c1 cases m %
157    ]
158  ]
159 qed. 
160
161 definition tuple_length ≝ λn.2*n+3.
162
163 axiom length_of_tuple: ∀n,t. tuple_TM n t → 
164   |t| = tuple_length n.
165
166 definition move_eq ≝ λm1,m2:move.
167   match m1 with
168   [R ⇒ match m2 with [R ⇒ true | _ ⇒ false]
169   |L ⇒ match m2 with [L ⇒ true | _ ⇒ false]
170   |N ⇒ match m2 with [N ⇒ true | _ ⇒ false]].
171   
172 definition tuples_of_pairs ≝ λn.λh.map … (λp.(tuple_of_pair n h p)@[〈bar,false〉]).
173
174 definition flatten ≝ λA.foldr (list A) (list A) (append A) [].
175
176 lemma wftable: ∀n,h,l.table_TM (S n) (flatten ? (tuples_of_pairs n h l)).
177 #n #h #l elim l // -l #a #tl #Hind 
178 whd in match (flatten … (tuples_of_pairs …));
179 >associative_append @ttm_cons //
180 qed.
181
182 lemma flatten_to_mem: ∀A,n,l,l1,l2.∀a:list A. 0 < n →
183   (∀x. mem ? x l → |x| = n) → |a| = n → flatten ? l = l1@a@l2  →
184     (∃q.|l1| = n*q)  → mem ? a l.
185 #A #n #l elim l
186   [normalize #l1 #l2 #a #posn #Hlen #Ha #Hnil @False_ind
187    cut (|a|=0) [@daemon] /2/
188   |#hd #tl #Hind #l1 #l2 #a #posn #Hlen #Ha 
189    whd in match (flatten ??); #Hflat * #q cases q
190     [<times_n_O #Hl1 
191      cut (a = hd) [@daemon] /2/
192     |#q1 #Hl1 lapply (split_exists … n l1 ?) //
193      * #l11 * #l12 * #Heql1 #Hlenl11 %2
194      @(Hind l12 l2 … posn ? Ha) 
195       [#x #memx @Hlen %2 //
196       |@(append_l2_injective ? hd l11) 
197         [>Hlenl11 @Hlen %1 %
198         |>Hflat >Heql1 >associative_append %
199         ]
200       |@(ex_intro …q1) @(injective_plus_r n) 
201        <Hlenl11 in ⊢ (??%?); <length_append <Heql1 >Hl1 //
202       ]
203     ]
204   ]
205 qed.
206
207 axiom match_decomp: ∀n,l,qin,cin,qout,cout,mv.
208   match_in_table (S n) qin cin qout cout mv l →
209   ∃l1,l2. l = l1@((mk_tuple qin cin qout cout mv)@[〈bar,false〉])@l2 ∧
210     (∃q.|l1| = (S (tuple_length (S n)))*q) ∧ tuple_TM (S n) (mk_tuple qin cin qout cout mv).
211 (*
212 lemma match_tech: ∀n,l,qin,cin,qout,cout,mv.
213   (∀t. mem ? t l → |t| = |mk_tuple qin cin qout cout mv|) →
214   match_in_table (S n) qin cin qout cout mv (flatten ? l) → 
215     ∃p. p = mk_tuple qin cin qout cout mv ∧ mem ? p l.
216 #n #l #qin #cin #qout #cout #mv #Hlen #Hmatch 
217 @(ex_intro … (mk_tuple qin cin qout cout mv)) % //
218 @flatten_to_mem *)
219
220 lemma match_to_tuple: ∀n,h,l,qin,cin,qout,cout,mv.
221   match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_of_pairs n h l)) → 
222     ∃p. p = mk_tuple qin cin qout cout mv ∧ mem ? (p@[〈bar,false〉]) (tuples_of_pairs n h l).
223 #n #h #l #qin #cin #qout #cout #mv #Hmatch 
224 @(ex_intro … (mk_tuple qin cin qout cout mv)) % //
225 cases (match_decomp … Hmatch) #l1 * #l2 * * #Hflat #Hlen #Htuple
226 @(flatten_to_mem … Hflat … Hlen)  
227   [// 
228   |@daemon
229   |>length_append >(length_of_tuple … Htuple) normalize //
230   ]
231 qed.
232
233 lemma mem_map: ∀A,B.∀f:A→B.∀l,b. 
234   mem ? b (map … f l) → ∃a. mem ? a l ∧ f a = b.
235 #A #B #f #l elim l 
236   [#b normalize @False_ind
237   |#a #tl #Hind #b normalize *
238     [#eqb @(ex_intro … a) /3/
239     |#memb cases (Hind … memb) #a * #mema #eqb
240      @(ex_intro … a) /3/
241     ]
242   ]
243 qed.
244
245 lemma match_to_pair: ∀n,h,l,qin,cin,qout,cout,mv.
246   match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_of_pairs n h l)) → 
247     ∃p. tuple_of_pair n h p = mk_tuple qin cin qout cout mv ∧ mem ? p l.
248 #n #h #l #qin #cin #qout #cout #mv #Hmatch 
249 cases (match_to_tuple … Hmatch)
250 #p * #eqp #memb 
251 cases(mem_map … (λp.(tuple_of_pair n h p)@[〈bar,false〉]) … memb)
252 #p1 * #Hmem #H @(ex_intro … p1) % /2/
253 qed.
254
255 (* turning DeqMove into a DeqSet *)
256 lemma move_eq_true:∀m1,m2.
257   move_eq m1 m2 = true ↔ m1 = m2.
258 *
259   [* normalize [% #_ % |2,3: % #H destruct ]
260   |* normalize [1,3: % #H destruct |% #_ % ]
261   |* normalize [1,2: % #H destruct |% #_ % ]
262 qed.
263
264 definition DeqMove ≝ mk_DeqSet move move_eq move_eq_true.
265
266 unification hint 0 ≔ ;
267     X ≟ DeqMove
268 (* ---------------------------------------- *) ⊢ 
269     move ≡ carr X.
270
271 unification hint  0 ≔ m1,m2; 
272     X ≟ DeqMove
273 (* ---------------------------------------- *) ⊢ 
274     move_eq m1 m2 ≡ eqb X m1 m2.
275
276 (* turning DeqMove into a FinSet *)
277 definition move_enum ≝ [L;R;N].
278
279 lemma move_enum_unique: uniqueb ? [L;R;N] = true.
280 // qed.
281
282 lemma move_enum_complete: ∀x:move. memb ? x [L;R;N] = true.
283 * // qed.
284
285 definition FinMove ≝ 
286   mk_FinSet DeqMove [L;R;N] move_enum_unique move_enum_complete.
287
288 unification hint  0 ≔ ; 
289     X ≟ FinMove
290 (* ---------------------------------------- *) ⊢ 
291     move ≡ FinSetcarr X.
292
293 definition trans_source ≝ λn.FinProd (initN n) (FinOption FinBool).
294 definition trans_target ≝ λn.FinProd (initN n) (FinOption (FinProd FinBool FinMove)).
295
296 lemma match_to_trans: 
297   ∀n.∀trans: trans_source n → trans_target n.
298   ∀h,qin,cin,qout,cout,mv.
299   match_in_table (S n) qin cin qout cout mv (flatten ? (tuples_of_pairs n h (graph_enum ?? trans))) → 
300   ∃s,t. tuple_of_pair n h 〈s,t〉 = mk_tuple qin cin qout cout mv 
301     ∧ trans s = t.
302 #n #trans #h #qin #cin #qout #cout #mv #Hmatch
303 cases (match_to_pair … Hmatch) -Hmatch * #s #t * #Heq #Hmem
304 @(ex_intro … s) @(ex_intro … t) % // @graph_enum_correct 
305 @mem_to_memb @Hmem 
306 qed.
307