]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/turing/multi_universal/alphabet.ma
more unistep
[helm.git] / matita / matita / lib / turing / multi_universal / alphabet.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basics/finset.ma".
16
17 inductive unialpha : Type[0] ≝ 
18 | bit : bool → unialpha
19 | bar : unialpha.
20
21 definition unialpha_eq ≝ 
22   λa1,a2.match a1 with
23   [ bit x ⇒ match a2 with [ bit y ⇒ ¬ xorb x y | _ ⇒ false ]
24   | bar ⇒ match a2 with [ bar ⇒ 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   | * [ #y ] normalize % #H1 destruct % ]
30 qed.
31
32 lemma unialpha_unique : 
33   uniqueb DeqUnialpha [bit true;bit false;bar] = true.
34 // qed.
35
36 lemma unialpha_complete :∀x:DeqUnialpha. 
37   memb ? x [bit true;bit false;bar] = true.
38 * // * //
39 qed.
40
41 definition FSUnialpha ≝ 
42   mk_FinSet DeqUnialpha [bit true;bit false;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.match c with [ bar ⇒ true | _ ⇒ false ].