]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/alphabet.ma
normal and tuples
[helm.git] / matita / matita / lib / turing / multi_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 include "basics/finset.ma".
13 include "basics/lists/list.ma".
14
15 inductive unialpha : Type[0] ≝ 
16 | bit  : bool → unialpha
17 | null : unialpha
18 | bar  : unialpha.
19
20 definition unialpha_eq ≝ 
21   λa1,a2.match a1 with
22   [ bit x ⇒ match a2 with [ bit y ⇒ ¬ xorb x y | _ ⇒ false ]
23   | bar ⇒ match a2 with [ bar ⇒ true | _ ⇒ false ] 
24   | null ⇒ match a2 with [ null ⇒ true | _ ⇒ false ] ].
25   
26 definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?.
27 * [ #x * [ #y cases x cases y normalize % // #Hfalse destruct
28          | *: normalize % #Hfalse destruct ]
29   | *: * [1,4: #y ] normalize % #H1 destruct % ]
30 qed.
31
32 lemma unialpha_unique : 
33   uniqueb DeqUnialpha [bit true;bit false;null;bar] = true.
34 // qed.
35
36 lemma unialpha_complete :∀x:DeqUnialpha. 
37   memb ? x [bit true;bit false;null;bar] = true.
38 * // * //
39 qed.
40
41 definition FSUnialpha ≝ 
42   mk_FinSet DeqUnialpha [bit true;bit false;null;bar] 
43   unialpha_unique unialpha_complete.
44
45 (*************************** testing characters *******************************)
46 definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ].
47 definition is_bar ≝ λc:DeqUnialpha. c == bar.
48 definition is_null ≝ λc:DeqUnialpha. c == null.
49
50 definition only_bits ≝ λl.
51   ∀c.mem ? c l → is_bit c = true.
52
53 definition no_bars ≝ λl.
54   ∀c.mem ? c l → is_bar c = false.