X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fturing%2Funiversal%2Falphabet.ma;h=4d87907d23130f1a6937a99ccd0d3df371bc9b31;hb=7d466423e4696ab52cca1068611ceb86c16136b8;hp=04701aeec180b8b868e7388f186cb173900bb2ef;hpb=67948c54e4e20ef530d2d2115c9a056275360012;p=helm.git diff --git a/matita/matita/lib/turing/universal/alphabet.ma b/matita/matita/lib/turing/universal/alphabet.ma index 04701aeec..4d87907d2 100644 --- a/matita/matita/lib/turing/universal/alphabet.ma +++ b/matita/matita/lib/turing/universal/alphabet.ma @@ -39,11 +39,20 @@ definition DeqUnialpha ≝ mk_DeqSet unialpha unialpha_eq ?. |*: * [1,6,11,16: #y ] normalize % #H1 destruct % ] qed. -axiom unialpha_unique : uniqueb DeqUnialpha [bit true;bit false;comma;bar;grid] = true. +lemma unialpha_unique : + uniqueb DeqUnialpha [bit true;bit false;null;comma;bar;grid] = true. +// qed. + +lemma unialpha_complete :∀x:DeqUnialpha. + memb ? x [bit true;bit false;null;comma;bar;grid] = true. +* // * // +qed. definition FSUnialpha ≝ - mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid] unialpha_unique. + mk_FinSet DeqUnialpha [bit true;bit false;null;comma;bar;grid] + unialpha_unique unialpha_complete. +(*************************** testing characters *******************************) definition is_bit ≝ λc.match c with [ bit _ ⇒ true | _ ⇒ false ]. definition is_null ≝ λc.match c with [ null ⇒ true | _ ⇒ false ]. @@ -55,3 +64,42 @@ definition is_bar ≝ λc.match c with [ bar ⇒ true | _ ⇒ false ]. definition is_comma ≝ λc.match c with [ comma ⇒ true | _ ⇒ false ]. definition bit_or_null ≝ λc.is_bit c ∨ is_null c. + +lemma bit_not_grid: ∀d. is_bit d = true → is_grid d = false. +* // normalize #H destruct +qed. + +lemma bit_or_null_not_grid: ∀d. bit_or_null d = true → is_grid d = false. +* // normalize #H destruct +qed. + +lemma bit_not_bar: ∀d. is_bit d = true → is_bar d = false. +* // normalize #H destruct +qed. + +lemma bit_or_null_not_bar: ∀d. bit_or_null d = true → is_bar d = false. +* // normalize #H destruct +qed. +(**************************** testing strings *********************************) +definition is_marked ≝ λalpha.λp:FinProd … alpha FinBool. + let 〈x,b〉 ≝ p in b. + +definition STape ≝ FinProd … FSUnialpha FinBool. + +definition only_bits ≝ λl. + ∀c.memb STape c l = true → is_bit (\fst c) = true. + +definition only_bits_or_nulls ≝ λl. + ∀c.memb STape c l = true → bit_or_null (\fst c) = true. + +definition no_grids ≝ λl. + ∀c.memb STape c l = true → is_grid (\fst c) = false. + +definition no_bars ≝ λl. + ∀c.memb STape c l = true → is_bar (\fst c) = false. + +definition no_marks ≝ λl. + ∀c.memb STape c l = true → is_marked ? c = false. + +definition bar_or_grid ≝ λc:STape.is_bar (\fst c) ∨ is_grid (\fst c). +