]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/universal/alphabet.ma
Porting to the new defintion of finset
[helm.git] / matita / matita / lib / turing / universal / alphabet.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 (* ALphabet of the universal machine *)
14
15 include "basics/finset.ma".
16
17 (*
18 include "turing/if_machine.ma".
19 include "turing/universal/tests.ma". *)
20
21 inductive unialpha : Type[0] ≝ 
22 | bit : bool → unialpha
23 | null : unialpha
24 | comma : unialpha
25 | bar : unialpha
26 | grid : unialpha.
27
28 definition unialpha_eq ≝ 
29   λa1,a2.match a1 with
30   [ bit x ⇒ match a2 with [ bit y ⇒ ¬ xorb x y | _ ⇒ false ]
31   | null ⇒ match a2 with [ null ⇒ true | _ ⇒ false ]
32   | comma ⇒ match a2 with [ comma ⇒ true | _ ⇒ false ]
33   | bar ⇒ match a2 with [ bar ⇒ true | _ ⇒ false ]
34   | grid ⇒ match a2 with [ grid ⇒ true | _ ⇒ false ] ].
35   
36 definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?.
37 * [ #x * [ #y cases x cases y normalize % // #Hfalse destruct
38          | *: normalize % #Hfalse destruct ]
39   |*: * [1,6,11,16: #y ] normalize % #H1 destruct % ]
40 qed.
41
42 lemma unialpha_unique : 
43   uniqueb DeqUnialpha [bit true;bit false;null;comma;bar;grid] = true.
44 // qed.
45
46 lemma unialpha_complete :∀x:DeqUnialpha. 
47   memb ? x [bit true;bit false;null;comma;bar;grid] = true.
48 * // * //
49 qed.
50
51 definition FSUnialpha ≝ 
52   mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid] 
53   unialpha_unique unialpha_complete.
54
55 definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ].
56
57 definition is_null ≝ λc.match c with [ null ⇒ true | _ ⇒ false ].
58
59 definition is_grid ≝ λc.match c with [ grid ⇒ true | _ ⇒ false ].
60
61 definition is_bar ≝ λc.match c with [ bar ⇒ true | _ ⇒ false ].
62
63 definition is_comma ≝ λc.match c with [ comma ⇒ true | _ ⇒ false ].
64
65 definition bit_or_null ≝ λc.is_bit c ∨ is_null c.