]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/tuples.ma
0fe4508ed63979d2d28011aac21efca38ef9b524
[helm.git] / matita / matita / lib / turing / multi_universal / 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 (****************************** table of tuples *******************************)
14 include "turing/multi_universal/normalTM.ma".
15
16 (* a well formed table is a list of tuples *) 
17  
18 definition table_TM ≝ λn,l,h.flatten ? (tuples_list n h l).
19
20 lemma table_TM_cons: ∀n,h,t,o. 
21   table_TM n (t::o) h = (tuple_encoding n h t)@(table_TM n o h).
22 // qed.
23
24 (************************** matching in a table *******************************)
25 lemma list_to_table: ∀n,l,h,tup. mem ? tup (tuples_list n h l) →
26   ∃ll,lr.table_TM n l h = ll@tup@lr.
27 #n #l #h #tup elim l 
28   [@False_ind
29   |#hd #tl #Hind *
30     [#Htup %{[]} %{(table_TM n tl h)} >Htup %
31     |#H cases (Hind H) #ll * #lr #H1
32      %{((tuple_encoding n h hd)@ll)} %{lr} 
33      >associative_append <H1 %
34     ]
35   ]
36 qed.
37
38 definition is_config : nat → list unialpha → Prop ≝  
39  λn,t.∃qin,cin.
40  only_bits qin ∧ cin ≠ bar ∧ |qin| = S n ∧ t = bar::qin@[cin].
41
42 lemma table_to_list: ∀n,l,h,c. is_config n c → 
43   (∃ll,lr.table_TM n l h = ll@c@lr) → 
44     ∃out,t. tuple_encoding n h t = (c@out) ∧ mem ? t l.
45 #n #l #h #c * #qin * #cin * * * #H1 #H2 #H3 #H4  
46  * #ll * #lr lapply ll -ll elim l
47   [>H4 #ll cases ll normalize [|#hd #tl ] #Habs destruct
48   |#t1 #othert #Hind #ll >table_TM_cons #Htuple
49    cut (S n < |ll@c@lr|)
50      [<Htuple >length_append >(length_of_tuple  … (is_tuple … ))
51       /2 by transitive_lt, le_n/] #Hsplit lapply Htuple -Htuple
52    cases (is_tuple … n h t1) #q1 * #c1 * #q2 * #c2 * #m 
53    * * * * * * * #Hq1 #Hq2 #Hc1 #Hc2 #Hm #Hlen1 #Hlen2 
54    whd in ⊢ (???%→?); #Ht1 
55    (* if ll is empty we match the first tuple t1, otherwise
56       we match inside othert *)
57    cases ll
58     [>H4 >Ht1 normalize in ⊢ (???%→?); 
59      >associative_append whd in ⊢ (??%?→?); #Heq destruct (Heq) -Heq
60      >associative_append in e0; #e0
61      lapply (append_l1_injective  … e0) [>H3 @Hlen1] #Heq1
62      lapply (append_l2_injective  … e0) [>H3 @Hlen1]
63      normalize in ⊢ (???%→?); whd in ⊢ (??%?→?); #Htemp 
64      lapply (cons_injective_l ????? Htemp) #Hc1
65      lapply (cons_injective_r ????? Htemp) -Htemp #Heq2
66      %{(q2@[c2;m])} %{t1} % 
67       [>Ht1 >Heq1 >Hc1 @eq_f >associative_append % 
68       |%1 %
69       ]
70     |(* ll not nil *)
71      #b #tl >Ht1 normalize in ⊢ (???%→?); 
72      whd in ⊢ (??%?→?); #Heq destruct (Heq) 
73      cases (compare_append … e0) #l *
74       [* cases l 
75         [#_ #Htab cases (Hind [ ] (sym_eq … Htab)) #out * #t * #Ht #Hmemt
76          %{out} %{t} % // %2 //
77         |(* this case is absurd *) 
78          #al #tll #Heq1 >H4 #Heq2 @False_ind 
79          lapply (cons_injective_l ? bar … Heq2) #Hbar <Hbar in Heq1; #Heq1
80          @(absurd (mem ? bar (q1@(c1::q2@[c2; m]))))
81           [>Heq1 @mem_append_l2 %1 //
82           |% #Hmembar cases (mem_append ???? Hmembar) -Hmembar
83             [#Hmembar lapply(Hq1 bar Hmembar) normalize #Habs destruct (Habs)
84             |* [#Habs @absurd //]
85              #Hmembar cases (mem_append ???? Hmembar) -Hmembar
86               [#Hmembar lapply(Hq2 bar Hmembar) normalize #Habs destruct (Habs)
87               |* [#Habs @absurd //] #Hmembar @(absurd ?? Hm) @sym_eq @mem_single //
88               ]
89             ]
90           ]
91         ]
92       |* #Htl #Htab cases (Hind … Htab) #out * #t * #Ht #Hmemt
93        %{out} %{t} % // %2 //
94       ] 
95     ]
96   ]
97 qed.
98