# HG changeset patch # User kleing # Date 1187562138 -7200 # Node ID e3a2b75b1cf91634f680e63f796ef59639688717 # Parent 76f7a8c6e8423943f948c07c204116f045a5015d boolean algebras as locales and numbers as types by Brian Huffman diff -r 76f7a8c6e842 -r e3a2b75b1cf9 CONTRIBUTORS --- a/CONTRIBUTORS Sun Aug 19 21:21:37 2007 +0200 +++ b/CONTRIBUTORS Mon Aug 20 00:22:18 2007 +0200 @@ -7,6 +7,9 @@ Contributions to Isabelle 2007 ------------------------------ +* August 2007: Brian Huffman, PSU + HOL/Library/Boolean_Algebra and HOL/Library/Numeral_Type + * June 2007: Amine Chaieb, TUM Semiring normalization and Groebner Bases diff -r 76f7a8c6e842 -r e3a2b75b1cf9 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sun Aug 19 21:21:37 2007 +0200 +++ b/src/HOL/IsaMakefile Mon Aug 20 00:22:18 2007 +0200 @@ -226,8 +226,10 @@ Library/SCT_Interpretation.thy \ Library/SCT_Implementation.thy Library/Size_Change_Termination.thy \ Library/SCT_Examples.thy Library/sct.ML \ - Library/Pure_term.thy Library/Eval.thy Library/Eval_Witness.thy Library/Pretty_Int.thy \ - Library/Pretty_Char.thy Library/Pretty_Char_chr.thy Library/Abstract_Rat.thy + Library/Pure_term.thy Library/Eval.thy Library/Eval_Witness.thy \ + Library/Pretty_Int.thy \ + Library/Pretty_Char.thy Library/Pretty_Char_chr.thy Library/Abstract_Rat.thy\ + Library/Numeral_Type.thy Library/Boolean_Algebra.thy @cd Library; $(ISATOOL) usedir $(OUT)/HOL Library diff -r 76f7a8c6e842 -r e3a2b75b1cf9 src/HOL/Library/Boolean_Algebra.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Boolean_Algebra.thy Mon Aug 20 00:22:18 2007 +0200 @@ -0,0 +1,276 @@ +(* + ID: $Id$ + Author: Brian Huffman + + Boolean algebras as locales. +*) + +header {* Boolean Algebras *} + +theory Boolean_Algebra +imports Main +begin + +locale boolean = + fixes conj :: "'a => 'a => 'a" (infixr "\" 70) + fixes disj :: "'a => 'a => 'a" (infixr "\" 65) + fixes compl :: "'a => 'a" ("\ _" [81] 80) + fixes zero :: "'a" ("\") + fixes one :: "'a" ("\") + assumes conj_assoc: "(x \ y) \ z = x \ (y \ z)" + assumes disj_assoc: "(x \ y) \ z = x \ (y \ z)" + assumes conj_commute: "x \ y = y \ x" + assumes disj_commute: "x \ y = y \ x" + assumes conj_disj_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" + assumes disj_conj_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" + assumes conj_one_right: "x \ \ = x" + assumes disj_zero_right: "x \ \ = x" + assumes conj_cancel_right: "x \ \ x = \" + assumes disj_cancel_right: "x \ \ x = \" +begin + +lemmas disj_ac = + disj_assoc disj_commute + mk_left_commute [of "disj", OF disj_assoc disj_commute] + +lemmas conj_ac = + conj_assoc conj_commute + mk_left_commute [of "conj", OF conj_assoc conj_commute] + +lemma dual: "boolean disj conj compl one zero" +apply (rule boolean.intro) +apply (rule disj_assoc) +apply (rule conj_assoc) +apply (rule disj_commute) +apply (rule conj_commute) +apply (rule disj_conj_distrib) +apply (rule conj_disj_distrib) +apply (rule disj_zero_right) +apply (rule conj_one_right) +apply (rule disj_cancel_right) +apply (rule conj_cancel_right) +done + +text {* Complement *} + +lemma complement_unique: + assumes 1: "a \ x = \" + assumes 2: "a \ x = \" + assumes 3: "a \ y = \" + assumes 4: "a \ y = \" + shows "x = y" +proof - + have "(a \ x) \ (x \ y) = (a \ y) \ (x \ y)" using 1 3 by simp + hence "(x \ a) \ (x \ y) = (y \ a) \ (y \ x)" using conj_commute by simp + hence "x \ (a \ y) = y \ (a \ x)" using conj_disj_distrib by simp + hence "x \ \ = y \ \" using 2 4 by simp + thus "x = y" using conj_one_right by simp +qed + +lemma compl_unique: "[| x \ y = \; x \ y = \ |] ==> \ x = y" +by (rule complement_unique [OF conj_cancel_right disj_cancel_right]) + +lemma double_compl [simp]: "\ (\ x) = x" +proof (rule compl_unique) + from conj_cancel_right show "\ x \ x = \" by (simp add: conj_commute) + from disj_cancel_right show "\ x \ x = \" by (simp add: disj_commute) +qed + +lemma compl_eq_compl_iff [simp]: "(\ x = \ y) = (x = y)" +by (rule inj_eq [OF inj_on_inverseI], rule double_compl) + +text {* Conjunction *} + +lemma conj_absorb: "x \ x = x" +proof - + have "x \ x = (x \ x) \ \" using disj_zero_right by simp + also have "... = (x \ x) \ (x \ \ x)" using conj_cancel_right by simp + also have "... = x \ (x \ \ x)" using conj_disj_distrib by simp + also have "... = x \ \" using disj_cancel_right by simp + also have "... = x" using conj_one_right by simp + finally show ?thesis . +qed + +lemma conj_zero_right [simp]: "x \ \ = \" +proof - + have "x \ \ = x \ (x \ \ x)" using conj_cancel_right by simp + also have "... = (x \ x) \ \ x" using conj_assoc by simp + also have "... = x \ \ x" using conj_absorb by simp + also have "... = \" using conj_cancel_right by simp + finally show ?thesis . +qed + +lemma compl_one [simp]: "\ \ = \" +by (rule compl_unique [OF conj_zero_right disj_zero_right]) + +lemma conj_zero_left [simp]: "\ \ x = \" +by (subst conj_commute) (rule conj_zero_right) + +lemma conj_one_left [simp]: "\ \ x = x" +by (subst conj_commute) (rule conj_one_right) + +lemma conj_cancel_left [simp]: "\ x \ x = \" +by (subst conj_commute) (rule conj_cancel_right) + +lemma conj_left_absorb [simp]: "x \ (x \ y) = x \ y" +by (simp add: conj_assoc [symmetric] conj_absorb) + +lemma conj_disj_distrib2: + "(y \ z) \ x = (y \ x) \ (z \ x)" +by (simp add: conj_commute conj_disj_distrib) + +lemmas conj_disj_distribs = + conj_disj_distrib conj_disj_distrib2 + +text {* Disjunction *} + +lemma disj_absorb [simp]: "x \ x = x" +by (rule boolean.conj_absorb [OF dual]) + +lemma disj_one_right [simp]: "x \ \ = \" +by (rule boolean.conj_zero_right [OF dual]) + +lemma compl_zero [simp]: "\ \ = \" +by (rule boolean.compl_one [OF dual]) + +lemma disj_zero_left [simp]: "\ \ x = x" +by (rule boolean.conj_one_left [OF dual]) + +lemma disj_one_left [simp]: "\ \ x = \" +by (rule boolean.conj_zero_left [OF dual]) + +lemma disj_cancel_left [simp]: "\ x \ x = \" +by (rule boolean.conj_cancel_left [OF dual]) + +lemma disj_left_absorb [simp]: "x \ (x \ y) = x \ y" +by (rule boolean.conj_left_absorb [OF dual]) + +lemma disj_conj_distrib2: + "(y \ z) \ x = (y \ x) \ (z \ x)" +by (rule boolean.conj_disj_distrib2 [OF dual]) + +lemmas disj_conj_distribs = + disj_conj_distrib disj_conj_distrib2 + +text {* De Morgan's Laws *} + +lemma de_Morgan_conj [simp]: "\ (x \ y) = \ x \ \ y" +proof (rule compl_unique) + have "(x \ y) \ (\ x \ \ y) = ((x \ y) \ \ x) \ ((x \ y) \ \ y)" + by (rule conj_disj_distrib) + also have "... = (y \ (x \ \ x)) \ (x \ (y \ \ y))" + by (simp add: conj_ac) + finally show "(x \ y) \ (\ x \ \ y) = \" + by (simp add: conj_cancel_right conj_zero_right disj_zero_right) +next + have "(x \ y) \ (\ x \ \ y) = (x \ (\ x \ \ y)) \ (y \ (\ x \ \ y))" + by (rule disj_conj_distrib2) + also have "... = (\ y \ (x \ \ x)) \ (\ x \ (y \ \ y))" + by (simp add: disj_ac) + finally show "(x \ y) \ (\ x \ \ y) = \" + by (simp add: disj_cancel_right disj_one_right conj_one_right) +qed + +lemma de_Morgan_disj [simp]: "\ (x \ y) = \ x \ \ y" +by (rule boolean.de_Morgan_conj [OF dual]) + +end + +text {* Symmetric Difference *} + +locale boolean_xor = boolean + + fixes xor :: "'a => 'a => 'a" (infixr "\" 65) + assumes xor_def: "x \ y = (x \ \ y) \ (\ x \ y)" +begin + +lemma xor_def2: + "x \ y = (x \ y) \ (\ x \ \ y)" +by (simp add: xor_def conj_disj_distribs + disj_ac conj_ac conj_cancel_right disj_zero_left) + +lemma xor_commute: "x \ y = y \ x" +by (simp add: xor_def conj_commute disj_commute) + +lemma xor_assoc: "(x \ y) \ z = x \ (y \ z)" +proof - + let ?t = "(x \ y \ z) \ (x \ \ y \ \ z) \ + (\ x \ y \ \ z) \ (\ x \ \ y \ z)" + have "?t \ (z \ x \ \ x) \ (z \ y \ \ y) = + ?t \ (x \ y \ \ y) \ (x \ z \ \ z)" + by (simp add: conj_cancel_right conj_zero_right) + thus "(x \ y) \ z = x \ (y \ z)" + apply (simp add: xor_def de_Morgan_disj de_Morgan_conj double_compl) + apply (simp add: conj_disj_distribs conj_ac disj_ac) + done +qed + +lemmas xor_ac = + xor_assoc xor_commute + mk_left_commute [of "xor", OF xor_assoc xor_commute] + +lemma xor_zero_right [simp]: "x \ \ = x" +by (simp add: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right) + +lemma xor_zero_left [simp]: "\ \ x = x" +by (subst xor_commute) (rule xor_zero_right) + +lemma xor_one_right [simp]: "x \ \ = \ x" +by (simp add: xor_def compl_one conj_zero_right conj_one_right disj_zero_left) + +lemma xor_one_left [simp]: "\ \ x = \ x" +by (subst xor_commute) (rule xor_one_right) + +lemma xor_self [simp]: "x \ x = \" +by (simp add: xor_def conj_cancel_right conj_cancel_left disj_zero_right) + +lemma xor_left_self [simp]: "x \ (x \ y) = y" +by (simp add: xor_assoc [symmetric] xor_self xor_zero_left) + +lemma xor_compl_left: "\ x \ y = \ (x \ y)" +apply (simp add: xor_def de_Morgan_disj de_Morgan_conj double_compl) +apply (simp add: conj_disj_distribs) +apply (simp add: conj_cancel_right conj_cancel_left) +apply (simp add: disj_zero_left disj_zero_right) +apply (simp add: disj_ac conj_ac) +done + +lemma xor_compl_right: "x \ \ y = \ (x \ y)" +apply (simp add: xor_def de_Morgan_disj de_Morgan_conj double_compl) +apply (simp add: conj_disj_distribs) +apply (simp add: conj_cancel_right conj_cancel_left) +apply (simp add: disj_zero_left disj_zero_right) +apply (simp add: disj_ac conj_ac) +done + +lemma xor_cancel_right [simp]: "x \ \ x = \" +by (simp add: xor_compl_right xor_self compl_zero) + +lemma xor_cancel_left [simp]: "\ x \ x = \" +by (subst xor_commute) (rule xor_cancel_right) + +lemma conj_xor_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" +proof - + have "(x \ y \ \ z) \ (x \ \ y \ z) = + (y \ x \ \ x) \ (z \ x \ \ x) \ (x \ y \ \ z) \ (x \ \ y \ z)" + by (simp add: conj_cancel_right conj_zero_right disj_zero_left) + thus "x \ (y \ z) = (x \ y) \ (x \ z)" + by (simp (no_asm_use) add: + xor_def de_Morgan_disj de_Morgan_conj double_compl + conj_disj_distribs conj_ac disj_ac) +qed + +lemma conj_xor_distrib2: + "(y \ z) \ x = (y \ x) \ (z \ x)" +proof - + have "x \ (y \ z) = (x \ y) \ (x \ z)" + by (rule conj_xor_distrib) + thus "(y \ z) \ x = (y \ x) \ (z \ x)" + by (simp add: conj_commute) +qed + +lemmas conj_xor_distribs = + conj_xor_distrib conj_xor_distrib2 + +end + +end diff -r 76f7a8c6e842 -r e3a2b75b1cf9 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Sun Aug 19 21:21:37 2007 +0200 +++ b/src/HOL/Library/Library.thy Mon Aug 20 00:22:18 2007 +0200 @@ -6,6 +6,7 @@ AssocList BigO Binomial + Boolean_Algebra Char_ord Coinductive_List Commutative_Ring @@ -24,6 +25,7 @@ NatPair Nat_Infinity Nested_Environment + Numeral_Type OptionalSugar Parity Permutation diff -r 76f7a8c6e842 -r e3a2b75b1cf9 src/HOL/Library/Numeral_Type.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Numeral_Type.thy Mon Aug 20 00:22:18 2007 +0200 @@ -0,0 +1,238 @@ +(* + ID: $Id$ + Author: Brian Huffman + + Numeral Syntax for Types +*) + +header "Numeral Syntax for Types" + +theory Numeral_Type + imports Infinite_Set +begin + +subsection {* Preliminary lemmas *} +(* These should be moved elsewhere *) + +lemma inj_Inl [simp]: "inj_on Inl A" + by (rule inj_onI, simp) + +lemma inj_Inr [simp]: "inj_on Inr A" + by (rule inj_onI, simp) + +lemma inj_Some [simp]: "inj_on Some A" + by (rule inj_onI, simp) + +lemma card_Plus: + "[| finite A; finite B |] ==> card (A <+> B) = card A + card B" + unfolding Plus_def + apply (subgoal_tac "Inl ` A \ Inr ` B = {}") + apply (simp add: card_Un_disjoint card_image) + apply fast + done + +lemma (in type_definition) univ: + "UNIV = Abs ` A" +proof + show "Abs ` A \ UNIV" by (rule subset_UNIV) + show "UNIV \ Abs ` A" + proof + fix x :: 'b + have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) + moreover have "Rep x \ A" by (rule Rep) + ultimately show "x \ Abs ` A" by (rule image_eqI) + qed +qed + +lemma (in type_definition) card: "card (UNIV :: 'b set) = card A" + by (simp add: univ card_image inj_on_def Abs_inject) + + +subsection {* Cardinalities of types *} + +syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))") + +translations "CARD(t)" => "card (UNIV::t set)" + +lemma card_unit: "CARD(unit) = 1" + unfolding univ_unit by simp + +lemma card_bool: "CARD(bool) = 2" + unfolding univ_bool by simp + +lemma card_prod: "CARD('a::finite \ 'b::finite) = CARD('a) * CARD('b)" + unfolding univ_prod by (simp only: card_cartesian_product) + +lemma card_sum: "CARD('a::finite + 'b::finite) = CARD('a) + CARD('b)" + unfolding univ_sum by (simp only: finite card_Plus) + +lemma card_option: "CARD('a::finite option) = Suc CARD('a)" + unfolding univ_option + apply (subgoal_tac "(None::'a option) \ range Some") + apply (simp add: finite card_image) + apply fast + done + +lemma card_set: "CARD('a::finite set) = 2 ^ CARD('a)" + unfolding univ_set + by (simp only: card_Pow finite numeral_2_eq_2) + +subsection {* Numeral Types *} + +typedef (open) pls = "UNIV :: nat set" .. +typedef (open) num1 = "UNIV :: unit set" .. +typedef (open) 'a bit0 = "UNIV :: (bool * 'a) set" .. +typedef (open) 'a bit1 = "UNIV :: (bool * 'a) option set" .. + +instance num1 :: finite +proof + show "finite (UNIV::num1 set)" + unfolding type_definition.univ [OF type_definition_num1] + using finite by (rule finite_imageI) +qed + +instance bit0 :: (finite) finite +proof + show "finite (UNIV::'a bit0 set)" + unfolding type_definition.univ [OF type_definition_bit0] + using finite by (rule finite_imageI) +qed + +instance bit1 :: (finite) finite +proof + show "finite (UNIV::'a bit1 set)" + unfolding type_definition.univ [OF type_definition_bit1] + using finite by (rule finite_imageI) +qed + +lemma card_num1: "CARD(num1) = 1" + unfolding type_definition.card [OF type_definition_num1] + by (simp only: card_unit) + +lemma card_bit0: "CARD('a::finite bit0) = 2 * CARD('a)" + unfolding type_definition.card [OF type_definition_bit0] + by (simp only: card_prod card_bool) + +lemma card_bit1: "CARD('a::finite bit1) = Suc (2 * CARD('a))" + unfolding type_definition.card [OF type_definition_bit1] + by (simp only: card_prod card_option card_bool) + +lemma card_pls: "CARD (pls) = 0" + by (simp add: type_definition.card [OF type_definition_pls]) + +lemmas card_univ_simps [simp] = + card_unit + card_bool + card_prod + card_sum + card_option + card_set + card_num1 + card_bit0 + card_bit1 + card_pls + +subsection {* Syntax *} + + +syntax + "_NumeralType" :: "num_const => type" ("_") + "_NumeralType0" :: type ("0") + "_NumeralType1" :: type ("1") + +translations + "_NumeralType1" == (type) "num1" + "_NumeralType0" == (type) "pls" + +parse_translation {* +let + +val num1_const = Syntax.const "Numeral_Type.num1"; +val pls_const = Syntax.const "Numeral_Type.pls"; +val B0_const = Syntax.const "Numeral_Type.bit0"; +val B1_const = Syntax.const "Numeral_Type.bit1"; + +fun mk_bintype n = + let + fun mk_bit n = if n = 0 then B0_const else B1_const; + fun bin_of n = + if n = 1 then num1_const + else if n = 0 then pls_const + else if n = ~1 then raise TERM ("negative type numeral", []) + else + let val (q, r) = IntInf.divMod (n, 2); + in mk_bit r $ bin_of q end; + in bin_of n end; + +fun numeral_tr (*"_NumeralType"*) [Const (str, _)] = + mk_bintype (valOf (IntInf.fromString str)) + | numeral_tr (*"_NumeralType"*) ts = raise TERM ("numeral_tr", ts); + +in [("_NumeralType", numeral_tr)] end; +*} + +print_translation {* +let +fun int_of [] = 0 + | int_of (b :: bs) = IntInf.fromInt b + (2 * int_of bs); + +fun bin_of (Const ("pls", _)) = [] + | bin_of (Const ("num1", _)) = [1] + | bin_of (Const ("bit0", _) $ bs) = 0 :: bin_of bs + | bin_of (Const ("bit1", _) $ bs) = 1 :: bin_of bs + | bin_of t = raise TERM("bin_of", [t]); + +fun bit_tr' b [t] = + let + val rev_digs = b :: bin_of t handle TERM _ => raise Match + val i = int_of rev_digs; + val num = IntInf.toString (IntInf.abs i); + in + Syntax.const "_NumeralType" $ Syntax.free num + end + | bit_tr' b _ = raise Match; + +in [("bit0", bit_tr' 0), ("bit1", bit_tr' 1)] end; +*} + + +subsection {* Classes with at values least 1 and 2 *} + +text {* Class finite already captures "at least 1" *} + +lemma zero_less_card_finite: + "0 < CARD('a::finite)" +proof (cases "CARD('a::finite) = 0") + case False thus ?thesis by (simp del: card_0_eq) +next + case True + thus ?thesis by (simp add: finite) +qed + +lemma one_le_card_finite: + "Suc 0 <= CARD('a::finite)" + by (simp add: less_Suc_eq_le [symmetric] zero_less_card_finite) + + +text {* Class for cardinality "at least 2" *} + +class card2 = finite + + assumes two_le_card: "2 <= CARD('a)" + +lemma one_less_card: "Suc 0 < CARD('a::card2)" + using two_le_card [where 'a='a] by simp + +instance bit0 :: (finite) card2 + by intro_classes (simp add: one_le_card_finite) + +instance bit1 :: (finite) card2 + by intro_classes (simp add: one_le_card_finite) + +subsection {* Examples *} + +term "TYPE(10)" + +lemma "CARD(0) = 0" by simp +lemma "CARD(17) = 17" by simp + +end