# HG changeset patch # User bulwahn # Date 1264170389 -3600 # Node ID dcd0fa5cc6d3e165903a20bc54271ad9d68b83ac # Parent 3b1957113753dfe3a6cec48e3ae9ae578454df5f# Parent e1b8f273640415617663d88ded5403cd6ca6d7dc merged diff -r 3b1957113753 -r dcd0fa5cc6d3 NEWS --- a/NEWS Fri Jan 22 11:42:28 2010 +0100 +++ b/NEWS Fri Jan 22 15:26:29 2010 +0100 @@ -12,6 +12,15 @@ *** HOL *** +* Various old-style primrec specifications in the HOL theories have been +replaced by new-style primrec, especially in theory List. The corresponding +constants now have authentic syntax. INCOMPATIBILITY. + +* Reorganized theory Multiset.thy: less duplication, less historical +organization of sections, conversion from associations lists to +multisets, rudimentary code generation. Use insert_DiffM2 [symmetric] +instead of elem_imp_eq_diff_union, if needed. INCOMPATIBILITY. + * Reorganized theory Sum_Type.thy; Inl and Inr now have authentic syntax. INCOMPATIBILITY. diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Code_Numeral.thy Fri Jan 22 15:26:29 2010 +0100 @@ -296,15 +296,14 @@ setup {* fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} - false false Code_Printer.str) ["SML", "Haskell"] + false Code_Printer.literal_naive_numeral) ["SML", "Haskell"] #> Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} - false true Code_Printer.str "OCaml" + false Code_Printer.literal_numeral "OCaml" #> Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral} - false false Code_Printer.str "Scala" + false Code_Printer.literal_naive_numeral "Scala" *} code_reserved SML Int int -code_reserved OCaml Big_int code_reserved Scala Int code_const "op + \ code_numeral \ code_numeral \ code_numeral" diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/IsaMakefile Fri Jan 22 15:26:29 2010 +0100 @@ -109,6 +109,7 @@ $(SRC)/Tools/Code/code_ml.ML \ $(SRC)/Tools/Code/code_preproc.ML \ $(SRC)/Tools/Code/code_printer.ML \ + $(SRC)/Tools/Code/code_scala.ML \ $(SRC)/Tools/Code/code_target.ML \ $(SRC)/Tools/Code/code_thingol.ML \ $(SRC)/Tools/Code_Generator.thy \ diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Library/Code_Integer.thy --- a/src/HOL/Library/Code_Integer.thy Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Library/Code_Integer.thy Fri Jan 22 15:26:29 2010 +0100 @@ -25,9 +25,9 @@ setup {* fold (Numeral.add_code @{const_name number_int_inst.number_of_int} - true true Code_Printer.str) ["SML", "OCaml", "Haskell", "Scala"] + true Code_Printer.literal_numeral) ["SML", "OCaml", "Haskell", "Scala"] #> Numeral.add_code @{const_name number_int_inst.number_of_int} - true true (fn s => (Pretty.block o map Code_Printer.str) ["BigInt", s]) "Scala" + true Code_Printer.literal_numeral "Scala" *} code_const "Int.Pls" and "Int.Min" and "Int.Bit0" and "Int.Bit1" @@ -88,7 +88,7 @@ code_const pdivmod (SML "IntInf.divMod/ (IntInf.abs _,/ IntInf.abs _)") (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)") - (Haskell "divMod/ (abs _)/ (abs _))") + (Haskell "divMod/ (abs _)/ (abs _)") (Scala "!(_.abs '/% _.abs)") code_const "eq_class.eq \ int \ int \ bool" @@ -113,10 +113,7 @@ (SML "IntInf.fromInt") (OCaml "_") (Haskell "toEnum") - (Scala "!BigInt(_)") - -code_reserved SML IntInf -code_reserved Scala BigInt + (Scala "!BigInt((_))") text {* Evaluation *} diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Library/Efficient_Nat.thy Fri Jan 22 15:26:29 2010 +0100 @@ -287,6 +287,8 @@ code_reserved Haskell Nat code_include Scala "Nat" {* +import scala.Math + object Nat { def apply(numeral: BigInt): Nat = new Nat(numeral max 0) @@ -309,7 +311,9 @@ def equals(that: Nat): Boolean = this.value == that.value def as_BigInt: BigInt = this.value - def as_Int: Int = this.value + def as_Int: Int = if (this.value >= Math.MAX_INT && this.value <= Math.MAX_INT) + this.value.intValue + else error("Int value too big:" + this.value.toString) def +(that: Nat): Nat = new Nat(this.value + that.value) def -(that: Nat): Nat = Nat(this.value + that.value) @@ -348,9 +352,9 @@ setup {* fold (Numeral.add_code @{const_name number_nat_inst.number_of_nat} - false true Code_Printer.str) ["SML", "OCaml", "Haskell"] + false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell"] #> Numeral.add_code @{const_name number_nat_inst.number_of_nat} - false true (fn s => (Pretty.block o map Code_Printer.str) ["Nat.Nat", s]) "Scala" + false Code_Printer.literal_positive_numeral "Scala" *} text {* diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Library/Multiset.thy Fri Jan 22 15:26:29 2010 +0100 @@ -2,34 +2,21 @@ Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker *) -header {* Multisets *} +header {* (Finite) multisets *} theory Multiset -imports List Main +imports Main begin subsection {* The type of multisets *} -typedef 'a multiset = "{f::'a => nat. finite {x . f x > 0}}" +typedef 'a multiset = "{f :: 'a => nat. finite {x. f x > 0}}" + morphisms count Abs_multiset proof show "(\x. 0::nat) \ ?multiset" by simp qed -lemmas multiset_typedef [simp] = - Abs_multiset_inverse Rep_multiset_inverse Rep_multiset - and [simp] = Rep_multiset_inject [symmetric] - -definition Mempty :: "'a multiset" ("{#}") where - [code del]: "{#} = Abs_multiset (\a. 0)" - -definition single :: "'a => 'a multiset" where - [code del]: "single a = Abs_multiset (\b. if b = a then 1 else 0)" - -definition count :: "'a multiset => 'a => nat" where - "count = Rep_multiset" - -definition MCollect :: "'a multiset => ('a => bool) => 'a multiset" where - "MCollect M P = Abs_multiset (\x. if P x then Rep_multiset M x else 0)" +lemmas multiset_typedef = Abs_multiset_inverse count_inverse count abbreviation Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50) where "a :# M == 0 < count M a" @@ -37,142 +24,456 @@ notation (xsymbols) Melem (infix "\#" 50) +lemma multiset_eq_conv_count_eq: + "M = N \ (\a. count M a = count N a)" + by (simp only: count_inject [symmetric] expand_fun_eq) + +lemma multi_count_ext: + "(\x. count A x = count B x) \ A = B" + using multiset_eq_conv_count_eq by auto + +text {* + \medskip Preservation of the representing set @{term multiset}. +*} + +lemma const0_in_multiset: + "(\a. 0) \ multiset" + by (simp add: multiset_def) + +lemma only1_in_multiset: + "(\b. if b = a then n else 0) \ multiset" + by (simp add: multiset_def) + +lemma union_preserves_multiset: + "M \ multiset \ N \ multiset \ (\a. M a + N a) \ multiset" + by (simp add: multiset_def) + +lemma diff_preserves_multiset: + assumes "M \ multiset" + shows "(\a. M a - N a) \ multiset" +proof - + have "{x. N x < M x} \ {x. 0 < M x}" + by auto + with assms show ?thesis + by (auto simp add: multiset_def intro: finite_subset) +qed + +lemma MCollect_preserves_multiset: + assumes "M \ multiset" + shows "(\x. if P x then M x else 0) \ multiset" +proof - + have "{x. (P x \ 0 < M x) \ P x} \ {x. 0 < M x}" + by auto + with assms show ?thesis + by (auto simp add: multiset_def intro: finite_subset) +qed + +lemmas in_multiset = const0_in_multiset only1_in_multiset + union_preserves_multiset diff_preserves_multiset MCollect_preserves_multiset + + +subsection {* Representing multisets *} + +text {* Multiset comprehension *} + +definition MCollect :: "'a multiset => ('a => bool) => 'a multiset" where + "MCollect M P = Abs_multiset (\x. if P x then count M x else 0)" + syntax "_MCollect" :: "pttrn => 'a multiset => bool => 'a multiset" ("(1{# _ :# _./ _#})") translations "{#x :# M. P#}" == "CONST MCollect M (\x. P)" -definition set_of :: "'a multiset => 'a set" where - "set_of M = {x. x :# M}" -instantiation multiset :: (type) "{plus, minus, zero, size}" +text {* Multiset enumeration *} + +instantiation multiset :: (type) "{zero, plus}" begin -definition union_def [code del]: - "M + N = Abs_multiset (\a. Rep_multiset M a + Rep_multiset N a)" - -definition diff_def [code del]: - "M - N = Abs_multiset (\a. Rep_multiset M a - Rep_multiset N a)" +definition Mempty_def: + "0 = Abs_multiset (\a. 0)" -definition Zero_multiset_def [simp]: - "0 = {#}" +abbreviation Mempty :: "'a multiset" ("{#}") where + "Mempty \ 0" -definition size_def: - "size M = setsum (count M) (set_of M)" +definition union_def: + "M + N = Abs_multiset (\a. count M a + count N a)" instance .. end -definition - multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where - "multiset_inter A B = A - (A - B)" +definition single :: "'a => 'a multiset" where + "single a = Abs_multiset (\b. if b = a then 1 else 0)" -text {* Multiset Enumeration *} syntax "_multiset" :: "args => 'a multiset" ("{#(_)#}") translations "{#x, xs#}" == "{#x#} + {#xs#}" "{#x#}" == "CONST single x" - -text {* - \medskip Preservation of the representing set @{term multiset}. -*} +lemma count_empty [simp]: "count {#} a = 0" + by (simp add: Mempty_def in_multiset multiset_typedef) -lemma const0_in_multiset: "(\a. 0) \ multiset" -by (simp add: multiset_def) - -lemma only1_in_multiset: "(\b. if b = a then 1 else 0) \ multiset" -by (simp add: multiset_def) - -lemma union_preserves_multiset: - "M \ multiset ==> N \ multiset ==> (\a. M a + N a) \ multiset" -by (simp add: multiset_def) +lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)" + by (simp add: single_def in_multiset multiset_typedef) -lemma diff_preserves_multiset: - "M \ multiset ==> (\a. M a - N a) \ multiset" -apply (simp add: multiset_def) -apply (rule finite_subset) - apply auto -done - -lemma MCollect_preserves_multiset: - "M \ multiset ==> (\x. if P x then M x else 0) \ multiset" -apply (simp add: multiset_def) -apply (rule finite_subset, auto) -done - -lemmas in_multiset = const0_in_multiset only1_in_multiset - union_preserves_multiset diff_preserves_multiset MCollect_preserves_multiset - - -subsection {* Algebraic properties *} +subsection {* Basic operations *} subsubsection {* Union *} -lemma union_empty [simp]: "M + {#} = M \ {#} + M = M" -by (simp add: union_def Mempty_def in_multiset) - -lemma union_commute: "M + N = N + (M::'a multiset)" -by (simp add: union_def add_ac in_multiset) - -lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" -by (simp add: union_def add_ac in_multiset) +lemma count_union [simp]: "count (M + N) a = count M a + count N a" + by (simp add: union_def in_multiset multiset_typedef) -lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" -proof - - have "M + (N + K) = (N + K) + M" by (rule union_commute) - also have "\ = N + (K + M)" by (rule union_assoc) - also have "K + M = M + K" by (rule union_commute) - finally show ?thesis . -qed - -lemmas union_ac = union_assoc union_commute union_lcomm - -instance multiset :: (type) comm_monoid_add -proof - fix a b c :: "'a multiset" - show "(a + b) + c = a + (b + c)" by (rule union_assoc) - show "a + b = b + a" by (rule union_commute) - show "0 + a = a" by simp -qed +instance multiset :: (type) cancel_comm_monoid_add proof +qed (simp_all add: multiset_eq_conv_count_eq) subsubsection {* Difference *} +instantiation multiset :: (type) minus +begin + +definition diff_def: + "M - N = Abs_multiset (\a. count M a - count N a)" + +instance .. + +end + +lemma count_diff [simp]: "count (M - N) a = count M a - count N a" + by (simp add: diff_def in_multiset multiset_typedef) + lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" -by (simp add: Mempty_def diff_def in_multiset) + by (simp add: Mempty_def diff_def in_multiset multiset_typedef) lemma diff_union_inverse2 [simp]: "M + {#a#} - {#a#} = M" -by (simp add: union_def diff_def in_multiset) + by (rule multi_count_ext) + (auto simp del: count_single simp add: union_def diff_def in_multiset multiset_typedef) lemma diff_cancel: "A - A = {#}" -by (simp add: diff_def Mempty_def) + by (rule multi_count_ext) simp + +lemma insert_DiffM: + "x \# M \ {#x#} + (M - {#x#}) = M" + by (clarsimp simp: multiset_eq_conv_count_eq) + +lemma insert_DiffM2 [simp]: + "x \# M \ M - {#x#} + {#x#} = M" + by (clarsimp simp: multiset_eq_conv_count_eq) + +lemma diff_right_commute: + "(M::'a multiset) - N - Q = M - Q - N" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma diff_union_swap: + "a \ b \ M - {#a#} + {#b#} = M + {#b#} - {#a#}" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma diff_union_single_conv: + "a \# J \ I + J - {#a#} = I + (J - {#a#})" + by (simp add: multiset_eq_conv_count_eq) -subsubsection {* Count of elements *} +subsubsection {* Intersection *} + +definition multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where + "multiset_inter A B = A - (A - B)" + +lemma multiset_inter_count: + "count (A #\ B) x = min (count A x) (count B x)" + by (simp add: multiset_inter_def multiset_typedef) -lemma count_empty [simp]: "count {#} a = 0" -by (simp add: count_def Mempty_def in_multiset) +lemma multiset_inter_commute: "A #\ B = B #\ A" + by (rule multi_count_ext) (simp add: multiset_inter_count) + +lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" + by (rule multi_count_ext) (simp add: multiset_inter_count) + +lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" + by (rule multi_count_ext) (simp add: multiset_inter_count) -lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)" -by (simp add: count_def single_def in_multiset) +lemmas multiset_inter_ac = + multiset_inter_commute + multiset_inter_assoc + multiset_inter_left_commute + +lemma multiset_inter_single: "a \ b \ {#a#} #\ {#b#} = {#}" + by (rule multi_count_ext) (auto simp add: multiset_inter_count) -lemma count_union [simp]: "count (M + N) a = count M a + count N a" -by (simp add: count_def union_def in_multiset) +lemma multiset_union_diff_commute: + assumes "B #\ C = {#}" + shows "A + B - C = A - C + B" +proof (rule multi_count_ext) + fix x + from assms have "min (count B x) (count C x) = 0" + by (auto simp add: multiset_inter_count multiset_eq_conv_count_eq) + then have "count B x = 0 \ count C x = 0" + by auto + then show "count (A + B - C) x = count (A - C + B) x" + by auto +qed -lemma count_diff [simp]: "count (M - N) a = count M a - count N a" -by (simp add: count_def diff_def in_multiset) + +subsubsection {* Comprehension (filter) *} lemma count_MCollect [simp]: "count {# x:#M. P x #} a = (if P a then count M a else 0)" -by (simp add: count_def MCollect_def in_multiset) + by (simp add: MCollect_def in_multiset multiset_typedef) + +lemma MCollect_empty [simp]: "MCollect {#} P = {#}" + by (rule multi_count_ext) simp + +lemma MCollect_single [simp]: + "MCollect {#x#} P = (if P x then {#x#} else {#})" + by (rule multi_count_ext) simp + +lemma MCollect_union [simp]: + "MCollect (M + N) f = MCollect M f + MCollect N f" + by (rule multi_count_ext) simp + + +subsubsection {* Equality of multisets *} + +lemma single_not_empty [simp]: "{#a#} \ {#} \ {#} \ {#a#}" + by (simp add: multiset_eq_conv_count_eq) + +lemma single_eq_single [simp]: "{#a#} = {#b#} \ a = b" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma union_eq_empty [iff]: "M + N = {#} \ M = {#} \ N = {#}" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma empty_eq_union [iff]: "{#} = M + N \ M = {#} \ N = {#}" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma multi_self_add_other_not_self [simp]: "M = M + {#x#} \ False" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma diff_single_trivial: + "\ x \# M \ M - {#x#} = M" + by (auto simp add: multiset_eq_conv_count_eq) + +lemma diff_single_eq_union: + "x \# M \ M - {#x#} = N \ M = N + {#x#}" + by auto + +lemma union_single_eq_diff: + "M + {#x#} = N \ M = N - {#x#}" + by (auto dest: sym) + +lemma union_single_eq_member: + "M + {#x#} = N \ x \# N" + by auto + +lemma union_is_single: + "M + N = {#a#} \ M = {#a#} \ N={#} \ M = {#} \ N = {#a#}" (is "?lhs = ?rhs") +proof + assume ?rhs then show ?lhs by auto +next + assume ?lhs + then have "\b. count (M + N) b = (if b = a then 1 else 0)" by auto + then have *: "\b. count M b + count N b = (if b = a then 1 else 0)" by auto + then have "count M a + count N a = 1" by auto + then have **: "count M a = 1 \ count N a = 0 \ count M a = 0 \ count N a = 1" + by auto + from * have "\b. b \ a \ count M b + count N b = 0" by auto + then have ***: "\b. b \ a \ count M b = 0 \ count N b = 0" by auto + from ** and *** have + "(\b. count M b = (if b = a then 1 else 0) \ count N b = 0) \ + (\b. count M b = 0 \ count N b = (if b = a then 1 else 0))" + by auto + then have + "(\b. count M b = (if b = a then 1 else 0)) \ (\b. count N b = 0) \ + (\b. count M b = 0) \ (\b. count N b = (if b = a then 1 else 0))" + by auto + then show ?rhs by (auto simp add: multiset_eq_conv_count_eq) +qed + +lemma single_is_union: + "{#a#} = M + N \ {#a#} = M \ N = {#} \ M = {#} \ {#a#} = N" + by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single) + +lemma add_eq_conv_diff: + "M + {#a#} = N + {#b#} \ M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#}" (is "?lhs = ?rhs") +proof + assume ?rhs then show ?lhs + by (auto simp add: add_assoc add_commute [of "{#b#}"]) + (drule sym, simp add: add_assoc [symmetric]) +next + assume ?lhs + show ?rhs + proof (cases "a = b") + case True with `?lhs` show ?thesis by simp + next + case False + from `?lhs` have "a \# N + {#b#}" by (rule union_single_eq_member) + with False have "a \# N" by auto + moreover from `?lhs` have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff) + moreover note False + ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"] diff_union_swap) + qed +qed + +lemma insert_noteq_member: + assumes BC: "B + {#b#} = C + {#c#}" + and bnotc: "b \ c" + shows "c \# B" +proof - + have "c \# C + {#c#}" by simp + have nc: "\ c \# {#b#}" using bnotc by simp + then have "c \# B + {#b#}" using BC by simp + then show "c \# B" using nc by simp +qed + +lemma add_eq_conv_ex: + "(M + {#a#} = N + {#b#}) = + (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" + by (auto simp add: add_eq_conv_diff) + + +subsubsection {* Pointwise ordering induced by count *} + +definition mset_le :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where + "A \# B \ (\a. count A a \ count B a)" + +definition mset_less :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where + "A <# B \ A \# B \ A \ B" + +notation mset_le (infix "\#" 50) +notation mset_less (infix "\#" 50) + +lemma mset_less_eqI: + "(\x. count A x \ count B x) \ A \# B" + by (simp add: mset_le_def) + +lemma mset_le_refl[simp]: "A \# A" +unfolding mset_le_def by auto + +lemma mset_le_trans: "A \# B \ B \# C \ A \# C" +unfolding mset_le_def by (fast intro: order_trans) + +lemma mset_le_antisym: "A \# B \ B \# A \ A = B" +apply (unfold mset_le_def) +apply (rule multiset_eq_conv_count_eq [THEN iffD2]) +apply (blast intro: order_antisym) +done + +lemma mset_le_exists_conv: "(A \# B) = (\C. B = A + C)" +apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) +apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) +done + +lemma mset_le_mono_add_right_cancel[simp]: "(A + C \# B + C) = (A \# B)" +unfolding mset_le_def by auto + +lemma mset_le_mono_add_left_cancel[simp]: "(C + A \# C + B) = (A \# B)" +unfolding mset_le_def by auto + +lemma mset_le_mono_add: "\ A \# B; C \# D \ \ A + C \# B + D" +apply (unfold mset_le_def) +apply auto +apply (erule_tac x = a in allE)+ +apply auto +done + +lemma mset_le_add_left[simp]: "A \# A + B" +unfolding mset_le_def by auto + +lemma mset_le_add_right[simp]: "B \# A + B" +unfolding mset_le_def by auto + +lemma mset_le_single: "a :# B \ {#a#} \# B" +by (simp add: mset_le_def) + +lemma multiset_diff_union_assoc: "C \# B \ A + B - C = A + (B - C)" +by (simp add: multiset_eq_conv_count_eq mset_le_def) + +lemma mset_le_multiset_union_diff_commute: +assumes "B \# A" +shows "A - B + C = A + C - B" +proof - + from mset_le_exists_conv [of "B" "A"] assms have "\D. A = B + D" .. + from this obtain D where "A = B + D" .. + then show ?thesis + apply simp + apply (subst add_commute) + apply (subst multiset_diff_union_assoc) + apply simp + apply (simp add: diff_cancel) + apply (subst add_assoc) + apply (subst add_commute [of "B" _]) + apply (subst multiset_diff_union_assoc) + apply simp + apply (simp add: diff_cancel) + done +qed + +interpretation mset_order: order "op \#" "op <#" +proof qed (auto intro: order.intro mset_le_refl mset_le_antisym + mset_le_trans simp: mset_less_def) + +interpretation mset_order_cancel_semigroup: + pordered_cancel_ab_semigroup_add "op +" "op \#" "op <#" +proof qed (erule mset_le_mono_add [OF mset_le_refl]) + +interpretation mset_order_semigroup_cancel: + pordered_ab_semigroup_add_imp_le "op +" "op \#" "op <#" +proof qed simp + +lemma mset_lessD: "A \# B \ x \# A \ x \# B" +apply (clarsimp simp: mset_le_def mset_less_def) +apply (erule_tac x=x in allE) +apply auto +done + +lemma mset_leD: "A \# B \ x \# A \ x \# B" +apply (clarsimp simp: mset_le_def mset_less_def) +apply (erule_tac x = x in allE) +apply auto +done + +lemma mset_less_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" +apply (rule conjI) + apply (simp add: mset_lessD) +apply (clarsimp simp: mset_le_def mset_less_def) +apply safe + apply (erule_tac x = a in allE) + apply (auto split: split_if_asm) +done + +lemma mset_le_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" +apply (rule conjI) + apply (simp add: mset_leD) +apply (force simp: mset_le_def mset_less_def split: split_if_asm) +done + +lemma mset_less_of_empty[simp]: "A \# {#} \ False" + by (auto simp add: mset_less_def mset_le_def multiset_eq_conv_count_eq) + +lemma multi_psub_of_add_self[simp]: "A \# A + {#x#}" +by (auto simp: mset_le_def mset_less_def) + +lemma multi_psub_self[simp]: "A \# A = False" +by (auto simp: mset_le_def mset_less_def) + +lemma mset_less_add_bothsides: + "T + {#x#} \# S + {#x#} \ T \# S" +by (auto simp: mset_le_def mset_less_def) + +lemma mset_less_empty_nonempty: "({#} \# S) = (S \ {#})" +by (auto simp: mset_le_def mset_less_def) + +lemma mset_less_diff_self: "c \# B \ B - {#c#} \# B" + by (auto simp: mset_le_def mset_less_def multiset_eq_conv_count_eq) subsubsection {* Set of elements *} +definition set_of :: "'a multiset => 'a set" where + "set_of M = {x. x :# M}" + lemma set_of_empty [simp]: "set_of {#} = {}" by (simp add: set_of_def) @@ -183,7 +484,7 @@ by (auto simp add: set_of_def) lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})" -by (auto simp: set_of_def Mempty_def in_multiset count_def expand_fun_eq [where f="Rep_multiset M"]) +by (auto simp add: set_of_def multiset_eq_conv_count_eq) lemma mem_set_of_iff [simp]: "(x \ set_of M) = (x :# M)" by (auto simp add: set_of_def) @@ -191,18 +492,28 @@ lemma set_of_MCollect [simp]: "set_of {# x:#M. P x #} = set_of M \ {x. P x}" by (auto simp add: set_of_def) +lemma finite_set_of [iff]: "finite (set_of M)" + using count [of M] by (simp add: multiset_def set_of_def) + subsubsection {* Size *} +instantiation multiset :: (type) size +begin + +definition size_def: + "size M = setsum (count M) (set_of M)" + +instance .. + +end + lemma size_empty [simp]: "size {#} = 0" by (simp add: size_def) lemma size_single [simp]: "size {#b#} = 1" by (simp add: size_def) -lemma finite_set_of [iff]: "finite (set_of M)" -using Rep_multiset [of M] by (simp add: multiset_def set_of_def count_def) - lemma setsum_count_Int: "finite A ==> setsum (count N) (A \ set_of N) = setsum (count N) A" apply (induct rule: finite_induct) @@ -221,9 +532,7 @@ done lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" -apply (unfold size_def Mempty_def count_def, auto simp: in_multiset) -apply (simp add: set_of_def count_def in_multiset expand_fun_eq) -done +by (auto simp add: size_def multiset_eq_conv_count_eq) lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) @@ -234,149 +543,16 @@ apply auto done - -subsubsection {* Equality of multisets *} - -lemma multiset_eq_conv_count_eq: "(M = N) = (\a. count M a = count N a)" -by (simp add: count_def expand_fun_eq) - -lemma single_not_empty [simp]: "{#a#} \ {#} \ {#} \ {#a#}" -by (simp add: single_def Mempty_def in_multiset expand_fun_eq) - -lemma single_eq_single [simp]: "({#a#} = {#b#}) = (a = b)" -by (auto simp add: single_def in_multiset expand_fun_eq) - -lemma union_eq_empty [iff]: "(M + N = {#}) = (M = {#} \ N = {#})" -by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) - -lemma empty_eq_union [iff]: "({#} = M + N) = (M = {#} \ N = {#})" -by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) - -lemma union_right_cancel [simp]: "(M + K = N + K) = (M = (N::'a multiset))" -by (simp add: union_def in_multiset expand_fun_eq) - -lemma union_left_cancel [simp]: "(K + M = K + N) = (M = (N::'a multiset))" -by (simp add: union_def in_multiset expand_fun_eq) - -lemma union_is_single: - "(M + N = {#a#}) = (M = {#a#} \ N={#} \ M = {#} \ N = {#a#})" -apply (simp add: Mempty_def single_def union_def in_multiset add_is_1 expand_fun_eq) -apply blast -done - -lemma single_is_union: - "({#a#} = M + N) \ ({#a#} = M \ N = {#} \ M = {#} \ {#a#} = N)" -apply (unfold Mempty_def single_def union_def) -apply (simp add: add_is_1 one_is_add in_multiset expand_fun_eq) -apply (blast dest: sym) -done - -lemma add_eq_conv_diff: - "(M + {#a#} = N + {#b#}) = - (M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#})" -using [[simproc del: neq]] -apply (unfold single_def union_def diff_def) -apply (simp (no_asm) add: in_multiset expand_fun_eq) -apply (rule conjI, force, safe, simp_all) -apply (simp add: eq_sym_conv) -done - -declare Rep_multiset_inject [symmetric, simp del] - -instance multiset :: (type) cancel_ab_semigroup_add -proof - fix a b c :: "'a multiset" - show "a + b = a + c \ b = c" by simp +lemma size_eq_Suc_imp_eq_union: + assumes "size M = Suc n" + shows "\a N. M = N + {#a#}" +proof - + from assms obtain a where "a \# M" + by (erule size_eq_Suc_imp_elem [THEN exE]) + then have "M = M - {#a#} + {#a#}" by simp + then show ?thesis by blast qed -lemma insert_DiffM: - "x \# M \ {#x#} + (M - {#x#}) = M" -by (clarsimp simp: multiset_eq_conv_count_eq) - -lemma insert_DiffM2[simp]: - "x \# M \ M - {#x#} + {#x#} = M" -by (clarsimp simp: multiset_eq_conv_count_eq) - -lemma multi_union_self_other_eq: - "(A::'a multiset) + X = A + Y \ X = Y" -by (induct A arbitrary: X Y) auto - -lemma multi_self_add_other_not_self[simp]: "(A = A + {#x#}) = False" -by (metis single_not_empty union_empty union_left_cancel) - -lemma insert_noteq_member: - assumes BC: "B + {#b#} = C + {#c#}" - and bnotc: "b \ c" - shows "c \# B" -proof - - have "c \# C + {#c#}" by simp - have nc: "\ c \# {#b#}" using bnotc by simp - then have "c \# B + {#b#}" using BC by simp - then show "c \# B" using nc by simp -qed - - -lemma add_eq_conv_ex: - "(M + {#a#} = N + {#b#}) = - (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" -by (auto simp add: add_eq_conv_diff) - - -lemma empty_multiset_count: - "(\x. count A x = 0) = (A = {#})" -by (metis count_empty multiset_eq_conv_count_eq) - - -subsubsection {* Intersection *} - -lemma multiset_inter_count: - "count (A #\ B) x = min (count A x) (count B x)" -by (simp add: multiset_inter_def) - -lemma multiset_inter_commute: "A #\ B = B #\ A" -by (simp add: multiset_eq_conv_count_eq multiset_inter_count - min_max.inf_commute) - -lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" -by (simp add: multiset_eq_conv_count_eq multiset_inter_count - min_max.inf_assoc) - -lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" -by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) - -lemmas multiset_inter_ac = - multiset_inter_commute - multiset_inter_assoc - multiset_inter_left_commute - -lemma multiset_inter_single: "a \ b \ {#a#} #\ {#b#} = {#}" -by (simp add: multiset_eq_conv_count_eq multiset_inter_count) - -lemma multiset_union_diff_commute: "B #\ C = {#} \ A + B - C = A - C + B" -apply (simp add: multiset_eq_conv_count_eq multiset_inter_count - split: split_if_asm) -apply clarsimp -apply (erule_tac x = a in allE) -apply auto -done - - -subsubsection {* Comprehension (filter) *} - -lemma MCollect_empty [simp]: "MCollect {#} P = {#}" -by (simp add: MCollect_def Mempty_def Abs_multiset_inject - in_multiset expand_fun_eq) - -lemma MCollect_single [simp]: - "MCollect {#x#} P = (if P x then {#x#} else {#})" -by (simp add: MCollect_def Mempty_def single_def Abs_multiset_inject - in_multiset expand_fun_eq) - -lemma MCollect_union [simp]: - "MCollect (M+N) f = MCollect M f + MCollect N f" -by (simp add: MCollect_def union_def Abs_multiset_inject - in_multiset expand_fun_eq) - subsection {* Induction and case splits *} @@ -434,17 +610,20 @@ shows "P M" proof - note defns = union_def single_def Mempty_def + note add' = add [unfolded defns, simplified] + have aux: "\a::'a. count (Abs_multiset (\b. if b = a then 1 else 0)) = + (\b. if b = a then 1 else 0)" by (simp add: Abs_multiset_inverse in_multiset) show ?thesis - apply (rule Rep_multiset_inverse [THEN subst]) - apply (rule Rep_multiset [THEN rep_multiset_induct]) + apply (rule count_inverse [THEN subst]) + apply (rule count [THEN rep_multiset_induct]) apply (rule empty [unfolded defns]) apply (subgoal_tac "f(b := f b + 1) = (\a. f a + (if a=b then 1 else 0))") prefer 2 apply (simp add: expand_fun_eq) apply (erule ssubst) apply (erule Abs_multiset_inverse [THEN subst]) - apply (drule add [unfolded defns, simplified]) - apply(simp add:in_multiset) + apply (drule add') + apply (simp add: aux) done qed @@ -470,18 +649,379 @@ apply (rule_tac x="M - {#x#}" in exI, simp) done +lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" +by (cases "B = {#}") (auto dest: multi_member_split) + lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \ P x #}" apply (subst multiset_eq_conv_count_eq) apply auto done -declare multiset_typedef [simp del] +lemma mset_less_size: "A \# B \ size A < size B" +proof (induct A arbitrary: B) + case (empty M) + then have "M \ {#}" by (simp add: mset_less_empty_nonempty) + then obtain M' x where "M = M' + {#x#}" + by (blast dest: multi_nonempty_split) + then show ?case by simp +next + case (add S x T) + have IH: "\B. S \# B \ size S < size B" by fact + have SxsubT: "S + {#x#} \# T" by fact + then have "x \# T" and "S \# T" by (auto dest: mset_less_insertD) + then obtain T' where T: "T = T' + {#x#}" + by (blast dest: multi_member_split) + then have "S \# T'" using SxsubT + by (blast intro: mset_less_add_bothsides) + then have "size S < size T'" using IH by simp + then show ?case using T by simp +qed + + +subsubsection {* Strong induction and subset induction for multisets *} + +text {* Well-foundedness of proper subset operator: *} + +text {* proper multiset subset *} + +definition + mset_less_rel :: "('a multiset * 'a multiset) set" where + "mset_less_rel = {(A,B). A \# B}" -lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" -by (cases "B = {#}") (auto dest: multi_member_split) +lemma multiset_add_sub_el_shuffle: + assumes "c \# B" and "b \ c" + shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" +proof - + from `c \# B` obtain A where B: "B = A + {#c#}" + by (blast dest: multi_member_split) + have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp + then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" + by (simp add: add_ac) + then show ?thesis using B by simp +qed + +lemma wf_mset_less_rel: "wf mset_less_rel" +apply (unfold mset_less_rel_def) +apply (rule wf_measure [THEN wf_subset, where f1=size]) +apply (clarsimp simp: measure_def inv_image_def mset_less_size) +done + +text {* The induction rules: *} + +lemma full_multiset_induct [case_names less]: +assumes ih: "\B. \A. A \# B \ P A \ P B" +shows "P B" +apply (rule wf_mset_less_rel [THEN wf_induct]) +apply (rule ih, auto simp: mset_less_rel_def) +done + +lemma multi_subset_induct [consumes 2, case_names empty add]: +assumes "F \# A" + and empty: "P {#}" + and insert: "\a F. a \# A \ P F \ P (F + {#a#})" +shows "P F" +proof - + from `F \# A` + show ?thesis + proof (induct F) + show "P {#}" by fact + next + fix x F + assume P: "F \# A \ P F" and i: "F + {#x#} \# A" + show "P (F + {#x#})" + proof (rule insert) + from i show "x \# A" by (auto dest: mset_le_insertD) + from i have "F \# A" by (auto dest: mset_le_insertD) + with P show "P F" . + qed + qed +qed -subsection {* Orderings *} +subsection {* Alternative representations *} + +subsubsection {* Lists *} + +primrec multiset_of :: "'a list \ 'a multiset" where + "multiset_of [] = {#}" | + "multiset_of (a # x) = multiset_of x + {# a #}" + +lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" +by (induct x) auto + +lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" +by (induct x) auto + +lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" +by (induct x) auto + +lemma mem_set_multiset_eq: "x \ set xs = (x :# multiset_of xs)" +by (induct xs) auto + +lemma multiset_of_append [simp]: + "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" + by (induct xs arbitrary: ys) (auto simp: add_ac) + +lemma surj_multiset_of: "surj multiset_of" +apply (unfold surj_def) +apply (rule allI) +apply (rule_tac M = y in multiset_induct) + apply auto +apply (rule_tac x = "x # xa" in exI) +apply auto +done + +lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}" +by (induct x) auto + +lemma distinct_count_atmost_1: + "distinct x = (! a. count (multiset_of x) a = (if a \ set x then 1 else 0))" +apply (induct x, simp, rule iffI, simp_all) +apply (rule conjI) +apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) +apply (erule_tac x = a in allE, simp, clarify) +apply (erule_tac x = aa in allE, simp) +done + +lemma multiset_of_eq_setD: + "multiset_of xs = multiset_of ys \ set xs = set ys" +by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) + +lemma set_eq_iff_multiset_of_eq_distinct: + "distinct x \ distinct y \ + (set x = set y) = (multiset_of x = multiset_of y)" +by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) + +lemma set_eq_iff_multiset_of_remdups_eq: + "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" +apply (rule iffI) +apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) +apply (drule distinct_remdups [THEN distinct_remdups + [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) +apply simp +done + +lemma multiset_of_compl_union [simp]: + "multiset_of [x\xs. P x] + multiset_of [x\xs. \P x] = multiset_of xs" + by (induct xs) (auto simp: add_ac) + +lemma count_filter: + "count (multiset_of xs) x = length [y \ xs. y = x]" +by (induct xs) auto + +lemma nth_mem_multiset_of: "i < length ls \ (ls ! i) :# multiset_of ls" +apply (induct ls arbitrary: i) + apply simp +apply (case_tac i) + apply auto +done + +lemma multiset_of_remove1: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}" +by (induct xs) (auto simp add: multiset_eq_conv_count_eq) + +lemma multiset_of_eq_length: +assumes "multiset_of xs = multiset_of ys" +shows "length xs = length ys" +using assms +proof (induct arbitrary: ys rule: length_induct) + case (1 xs ys) + show ?case + proof (cases xs) + case Nil with "1.prems" show ?thesis by simp + next + case (Cons x xs') + note xCons = Cons + show ?thesis + proof (cases ys) + case Nil + with "1.prems" Cons show ?thesis by simp + next + case (Cons y ys') + have x_in_ys: "x = y \ x \ set ys'" + proof (cases "x = y") + case True then show ?thesis .. + next + case False + from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp + with False show ?thesis by (simp add: mem_set_multiset_eq) + qed + from "1.hyps" have IH: "length xs' < length xs \ + (\x. multiset_of xs' = multiset_of x \ length xs' = length x)" by blast + from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))" + apply - + apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff) + apply fastsimp + done + with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp + from x_in_ys have "x \ y \ length ys' > 0" by auto + with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1) + qed + qed +qed + +text {* + This lemma shows which properties suffice to show that a function + @{text "f"} with @{text "f xs = ys"} behaves like sort. +*} +lemma properties_for_sort: + "multiset_of ys = multiset_of xs \ sorted ys \ sort xs = ys" +proof (induct xs arbitrary: ys) + case Nil then show ?case by simp +next + case (Cons x xs) + then have "x \ set ys" + by (auto simp add: mem_set_multiset_eq intro!: ccontr) + with Cons.prems Cons.hyps [of "remove1 x ys"] show ?case + by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1) +qed + +lemma multiset_of_remdups_le: "multiset_of (remdups xs) \# multiset_of xs" +apply (induct xs) + apply auto +apply (rule mset_le_trans) + apply auto +done + +lemma multiset_of_update: + "i < length ls \ multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}" +proof (induct ls arbitrary: i) + case Nil then show ?case by simp +next + case (Cons x xs) + show ?case + proof (cases i) + case 0 then show ?thesis by simp + next + case (Suc i') + with Cons show ?thesis + apply simp + apply (subst add_assoc) + apply (subst add_commute [of "{#v#}" "{#x#}"]) + apply (subst add_assoc [symmetric]) + apply simp + apply (rule mset_le_multiset_union_diff_commute) + apply (simp add: mset_le_single nth_mem_multiset_of) + done + qed +qed + +lemma multiset_of_swap: + "i < length ls \ j < length ls \ + multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" + by (cases "i = j") (simp_all add: multiset_of_update nth_mem_multiset_of) + + +subsubsection {* Association lists -- including rudimentary code generation *} + +definition count_of :: "('a \ nat) list \ 'a \ nat" where + "count_of xs x = (case map_of xs x of None \ 0 | Some n \ n)" + +lemma count_of_multiset: + "count_of xs \ multiset" +proof - + let ?A = "{x::'a. 0 < (case map_of xs x of None \ 0\nat | Some (n\nat) \ n)}" + have "?A \ dom (map_of xs)" + proof + fix x + assume "x \ ?A" + then have "0 < (case map_of xs x of None \ 0\nat | Some (n\nat) \ n)" by simp + then have "map_of xs x \ None" by (cases "map_of xs x") auto + then show "x \ dom (map_of xs)" by auto + qed + with finite_dom_map_of [of xs] have "finite ?A" + by (auto intro: finite_subset) + then show ?thesis + by (simp add: count_of_def expand_fun_eq multiset_def) +qed + +lemma count_simps [simp]: + "count_of [] = (\_. 0)" + "count_of ((x, n) # xs) = (\y. if x = y then n else count_of xs y)" + by (simp_all add: count_of_def expand_fun_eq) + +lemma count_of_empty: + "x \ fst ` set xs \ count_of xs x = 0" + by (induct xs) (simp_all add: count_of_def) + +lemma count_of_filter: + "count_of (filter (P \ fst) xs) x = (if P x then count_of xs x else 0)" + by (induct xs) auto + +definition Bag :: "('a \ nat) list \ 'a multiset" where + "Bag xs = Abs_multiset (count_of xs)" + +code_datatype Bag + +lemma count_Bag [simp, code]: + "count (Bag xs) = count_of xs" + by (simp add: Bag_def count_of_multiset Abs_multiset_inverse) + +lemma Mempty_Bag [code]: + "{#} = Bag []" + by (simp add: multiset_eq_conv_count_eq) + +lemma single_Bag [code]: + "{#x#} = Bag [(x, 1)]" + by (simp add: multiset_eq_conv_count_eq) + +lemma MCollect_Bag [code]: + "MCollect (Bag xs) P = Bag (filter (P \ fst) xs)" + by (simp add: multiset_eq_conv_count_eq count_of_filter) + +lemma mset_less_eq_Bag [code]: + "Bag xs \# A \ (\(x, n) \ set xs. count_of xs x \ count A x)" + (is "?lhs \ ?rhs") +proof + assume ?lhs then show ?rhs + by (auto simp add: mset_le_def count_Bag) +next + assume ?rhs + show ?lhs + proof (rule mset_less_eqI) + fix x + from `?rhs` have "count_of xs x \ count A x" + by (cases "x \ fst ` set xs") (auto simp add: count_of_empty) + then show "count (Bag xs) x \ count A x" + by (simp add: mset_le_def count_Bag) + qed +qed + +instantiation multiset :: (eq) eq +begin + +definition + "HOL.eq A B \ A \# B \ B \# A" + +instance proof +qed (simp add: eq_multiset_def mset_order.eq_iff) + +end + +definition (in term_syntax) + bagify :: "('a\typerep \ nat) list \ (unit \ Code_Evaluation.term) + \ 'a multiset \ (unit \ Code_Evaluation.term)" where + [code_unfold]: "bagify xs = Code_Evaluation.valtermify Bag {\} xs" + +notation fcomp (infixl "o>" 60) +notation scomp (infixl "o\" 60) + +instantiation multiset :: (random) random +begin + +definition + "Quickcheck.random i = Quickcheck.random i o\ (\xs. Pair (bagify xs))" + +instance .. + +end + +no_notation fcomp (infixl "o>" 60) +no_notation scomp (infixl "o\" 60) + +hide (open) const bagify + + +subsection {* The multiset order *} subsubsection {* Well-foundedness *} @@ -490,7 +1030,7 @@ (\b. b :# K --> (b, a) \ r)}" definition mult :: "('a \ 'a) set => ('a multiset \ 'a multiset) set" where - "mult r = (mult1 r)\<^sup>+" + [code del]: "mult r = (mult1 r)\<^sup>+" lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) @@ -523,7 +1063,7 @@ next fix K' assume "M0' = K' + {#a#}" - with N have n: "N = K' + K + {#a#}" by (simp add: union_ac) + with N have n: "N = K' + K + {#a#}" by (simp add: add_ac) assume "M0 = K' + {#a'#}" with r have "?R (K' + K) M0" by blast @@ -568,7 +1108,7 @@ with wf_hyp have "\M \ ?W. M + {#x#} \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "(M0 + K) + {#x#} \ ?W" .. - then show "M0 + (K + {#x#}) \ ?W" by (simp only: union_assoc) + then show "M0 + (K + {#x#}) \ ?W" by (simp only: add_assoc) qed then show "N \ ?W" by (simp only: N) qed @@ -610,11 +1150,6 @@ subsubsection {* Closure-free presentation *} -(*Badly needed: a linear arithmetic procedure for multisets*) - -lemma diff_union_single_conv: "a :# J ==> I + J - {#a#} = I + (J - {#a#})" -by (simp add: multiset_eq_conv_count_eq) - text {* One direction. *} lemma mult_implies_one_step: @@ -628,7 +1163,7 @@ apply (rule_tac x = I in exI) apply (simp (no_asm)) apply (rule_tac x = "(K - {#a#}) + Ka" in exI) - apply (simp (no_asm_simp) add: union_assoc [symmetric]) + apply (simp (no_asm_simp) add: add_assoc [symmetric]) apply (drule_tac f = "\M. M - {#a#}" in arg_cong) apply (simp add: diff_union_single_conv) apply (simp (no_asm_use) add: trans_def) @@ -649,14 +1184,6 @@ apply (simp (no_asm)) done -lemma elem_imp_eq_diff_union: "a :# M ==> M = M - {#a#} + {#a#}" -by (simp add: multiset_eq_conv_count_eq) - -lemma size_eq_Suc_imp_eq_union: "size M = Suc n ==> \a N. M = N + {#a#}" -apply (erule size_eq_Suc_imp_elem [THEN exE]) -apply (drule elem_imp_eq_diff_union, auto) -done - lemma one_step_implies_mult_aux: "trans r ==> \I J K. (size J = n \ J \ {#} \ (\k \ set_of K. \j \ set_of J. (k, j) \ r)) @@ -679,13 +1206,13 @@ (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") prefer 2 apply force -apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def) +apply (simp (no_asm_use) add: add_assoc [symmetric] mult_def) apply (erule trancl_trans) apply (rule r_into_trancl) apply (simp add: mult1_def set_of_def) apply (rule_tac x = a in exI) apply (rule_tac x = "I + J'" in exI) -apply (simp add: union_ac) +apply (simp add: add_ac) done lemma one_step_implies_mult: @@ -699,10 +1226,10 @@ instantiation multiset :: (order) order begin -definition less_multiset_def [code del]: +definition less_multiset_def: "M' < M \ (M', M) \ mult {(x', x). x' < x}" -definition le_multiset_def [code del]: +definition le_multiset_def: "M' <= M \ M' = M \ M' < (M::'a multiset)" lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}" @@ -776,7 +1303,7 @@ apply auto apply (rule_tac x = a in exI) apply (rule_tac x = "C + M0" in exI) -apply (simp add: union_assoc) +apply (simp add: add_assoc) done lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)" @@ -787,8 +1314,8 @@ done lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)" -apply (subst union_commute [of B C]) -apply (subst union_commute [of D C]) +apply (subst add_commute [of B C]) +apply (subst add_commute [of D C]) apply (erule union_less_mono2) done @@ -819,7 +1346,7 @@ qed lemma union_upper2: "B <= A + (B::'a::order multiset)" -by (subst union_commute) (rule union_upper1) +by (subst add_commute) (rule union_upper1) instance multiset :: (order) pordered_ab_semigroup_add apply intro_classes @@ -827,416 +1354,6 @@ done -subsection {* Link with lists *} - -primrec multiset_of :: "'a list \ 'a multiset" where - "multiset_of [] = {#}" | - "multiset_of (a # x) = multiset_of x + {# a #}" - -lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" -by (induct x) auto - -lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" -by (induct x) auto - -lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" -by (induct x) auto - -lemma mem_set_multiset_eq: "x \ set xs = (x :# multiset_of xs)" -by (induct xs) auto - -lemma multiset_of_append [simp]: - "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" -by (induct xs arbitrary: ys) (auto simp: union_ac) - -lemma surj_multiset_of: "surj multiset_of" -apply (unfold surj_def) -apply (rule allI) -apply (rule_tac M = y in multiset_induct) - apply auto -apply (rule_tac x = "x # xa" in exI) -apply auto -done - -lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}" -by (induct x) auto - -lemma distinct_count_atmost_1: - "distinct x = (! a. count (multiset_of x) a = (if a \ set x then 1 else 0))" -apply (induct x, simp, rule iffI, simp_all) -apply (rule conjI) -apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) -apply (erule_tac x = a in allE, simp, clarify) -apply (erule_tac x = aa in allE, simp) -done - -lemma multiset_of_eq_setD: - "multiset_of xs = multiset_of ys \ set xs = set ys" -by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) - -lemma set_eq_iff_multiset_of_eq_distinct: - "distinct x \ distinct y \ - (set x = set y) = (multiset_of x = multiset_of y)" -by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) - -lemma set_eq_iff_multiset_of_remdups_eq: - "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" -apply (rule iffI) -apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) -apply (drule distinct_remdups [THEN distinct_remdups - [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) -apply simp -done - -lemma multiset_of_compl_union [simp]: - "multiset_of [x\xs. P x] + multiset_of [x\xs. \P x] = multiset_of xs" -by (induct xs) (auto simp: union_ac) - -lemma count_filter: - "count (multiset_of xs) x = length [y \ xs. y = x]" -by (induct xs) auto - -lemma nth_mem_multiset_of: "i < length ls \ (ls ! i) :# multiset_of ls" -apply (induct ls arbitrary: i) - apply simp -apply (case_tac i) - apply auto -done - -lemma multiset_of_remove1: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}" -by (induct xs) (auto simp add: multiset_eq_conv_count_eq) - -lemma multiset_of_eq_length: -assumes "multiset_of xs = multiset_of ys" -shows "length xs = length ys" -using assms -proof (induct arbitrary: ys rule: length_induct) - case (1 xs ys) - show ?case - proof (cases xs) - case Nil with "1.prems" show ?thesis by simp - next - case (Cons x xs') - note xCons = Cons - show ?thesis - proof (cases ys) - case Nil - with "1.prems" Cons show ?thesis by simp - next - case (Cons y ys') - have x_in_ys: "x = y \ x \ set ys'" - proof (cases "x = y") - case True then show ?thesis .. - next - case False - from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp - with False show ?thesis by (simp add: mem_set_multiset_eq) - qed - from "1.hyps" have IH: "length xs' < length xs \ - (\x. multiset_of xs' = multiset_of x \ length xs' = length x)" by blast - from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))" - apply - - apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff) - apply fastsimp - done - with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp - from x_in_ys have "x \ y \ length ys' > 0" by auto - with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1) - qed - qed -qed - -text {* - This lemma shows which properties suffice to show that a function - @{text "f"} with @{text "f xs = ys"} behaves like sort. -*} -lemma properties_for_sort: - "multiset_of ys = multiset_of xs \ sorted ys \ sort xs = ys" -proof (induct xs arbitrary: ys) - case Nil then show ?case by simp -next - case (Cons x xs) - then have "x \ set ys" - by (auto simp add: mem_set_multiset_eq intro!: ccontr) - with Cons.prems Cons.hyps [of "remove1 x ys"] show ?case - by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1) -qed - - -subsection {* Pointwise ordering induced by count *} - -definition mset_le :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where - [code del]: "A \# B \ (\a. count A a \ count B a)" - -definition mset_less :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where - [code del]: "A <# B \ A \# B \ A \ B" - -notation mset_le (infix "\#" 50) -notation mset_less (infix "\#" 50) - -lemma mset_le_refl[simp]: "A \# A" -unfolding mset_le_def by auto - -lemma mset_le_trans: "A \# B \ B \# C \ A \# C" -unfolding mset_le_def by (fast intro: order_trans) - -lemma mset_le_antisym: "A \# B \ B \# A \ A = B" -apply (unfold mset_le_def) -apply (rule multiset_eq_conv_count_eq [THEN iffD2]) -apply (blast intro: order_antisym) -done - -lemma mset_le_exists_conv: "(A \# B) = (\C. B = A + C)" -apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) -apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) -done - -lemma mset_le_mono_add_right_cancel[simp]: "(A + C \# B + C) = (A \# B)" -unfolding mset_le_def by auto - -lemma mset_le_mono_add_left_cancel[simp]: "(C + A \# C + B) = (A \# B)" -unfolding mset_le_def by auto - -lemma mset_le_mono_add: "\ A \# B; C \# D \ \ A + C \# B + D" -apply (unfold mset_le_def) -apply auto -apply (erule_tac x = a in allE)+ -apply auto -done - -lemma mset_le_add_left[simp]: "A \# A + B" -unfolding mset_le_def by auto - -lemma mset_le_add_right[simp]: "B \# A + B" -unfolding mset_le_def by auto - -lemma mset_le_single: "a :# B \ {#a#} \# B" -by (simp add: mset_le_def) - -lemma multiset_diff_union_assoc: "C \# B \ A + B - C = A + (B - C)" -by (simp add: multiset_eq_conv_count_eq mset_le_def) - -lemma mset_le_multiset_union_diff_commute: -assumes "B \# A" -shows "A - B + C = A + C - B" -proof - - from mset_le_exists_conv [of "B" "A"] assms have "\D. A = B + D" .. - from this obtain D where "A = B + D" .. - then show ?thesis - apply simp - apply (subst union_commute) - apply (subst multiset_diff_union_assoc) - apply simp - apply (simp add: diff_cancel) - apply (subst union_assoc) - apply (subst union_commute[of "B" _]) - apply (subst multiset_diff_union_assoc) - apply simp - apply (simp add: diff_cancel) - done -qed - -lemma multiset_of_remdups_le: "multiset_of (remdups xs) \# multiset_of xs" -apply (induct xs) - apply auto -apply (rule mset_le_trans) - apply auto -done - -lemma multiset_of_update: - "i < length ls \ multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}" -proof (induct ls arbitrary: i) - case Nil then show ?case by simp -next - case (Cons x xs) - show ?case - proof (cases i) - case 0 then show ?thesis by simp - next - case (Suc i') - with Cons show ?thesis - apply simp - apply (subst union_assoc) - apply (subst union_commute [where M = "{#v#}" and N = "{#x#}"]) - apply (subst union_assoc [symmetric]) - apply simp - apply (rule mset_le_multiset_union_diff_commute) - apply (simp add: mset_le_single nth_mem_multiset_of) - done - qed -qed - -lemma multiset_of_swap: - "i < length ls \ j < length ls \ - multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" -apply (case_tac "i = j") - apply simp -apply (simp add: multiset_of_update) -apply (subst elem_imp_eq_diff_union[symmetric]) - apply (simp add: nth_mem_multiset_of) -apply simp -done - -interpretation mset_order: order "op \#" "op <#" -proof qed (auto intro: order.intro mset_le_refl mset_le_antisym - mset_le_trans simp: mset_less_def) - -interpretation mset_order_cancel_semigroup: - pordered_cancel_ab_semigroup_add "op +" "op \#" "op <#" -proof qed (erule mset_le_mono_add [OF mset_le_refl]) - -interpretation mset_order_semigroup_cancel: - pordered_ab_semigroup_add_imp_le "op +" "op \#" "op <#" -proof qed simp - - -lemma mset_lessD: "A \# B \ x \# A \ x \# B" -apply (clarsimp simp: mset_le_def mset_less_def) -apply (erule_tac x=x in allE) -apply auto -done - -lemma mset_leD: "A \# B \ x \# A \ x \# B" -apply (clarsimp simp: mset_le_def mset_less_def) -apply (erule_tac x = x in allE) -apply auto -done - -lemma mset_less_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" -apply (rule conjI) - apply (simp add: mset_lessD) -apply (clarsimp simp: mset_le_def mset_less_def) -apply safe - apply (erule_tac x = a in allE) - apply (auto split: split_if_asm) -done - -lemma mset_le_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" -apply (rule conjI) - apply (simp add: mset_leD) -apply (force simp: mset_le_def mset_less_def split: split_if_asm) -done - -lemma mset_less_of_empty[simp]: "A \# {#} = False" -by (induct A) (auto simp: mset_le_def mset_less_def) - -lemma multi_psub_of_add_self[simp]: "A \# A + {#x#}" -by (auto simp: mset_le_def mset_less_def) - -lemma multi_psub_self[simp]: "A \# A = False" -by (auto simp: mset_le_def mset_less_def) - -lemma mset_less_add_bothsides: - "T + {#x#} \# S + {#x#} \ T \# S" -by (auto simp: mset_le_def mset_less_def) - -lemma mset_less_empty_nonempty: "({#} \# S) = (S \ {#})" -by (auto simp: mset_le_def mset_less_def) - -lemma mset_less_size: "A \# B \ size A < size B" -proof (induct A arbitrary: B) - case (empty M) - then have "M \ {#}" by (simp add: mset_less_empty_nonempty) - then obtain M' x where "M = M' + {#x#}" - by (blast dest: multi_nonempty_split) - then show ?case by simp -next - case (add S x T) - have IH: "\B. S \# B \ size S < size B" by fact - have SxsubT: "S + {#x#} \# T" by fact - then have "x \# T" and "S \# T" by (auto dest: mset_less_insertD) - then obtain T' where T: "T = T' + {#x#}" - by (blast dest: multi_member_split) - then have "S \# T'" using SxsubT - by (blast intro: mset_less_add_bothsides) - then have "size S < size T'" using IH by simp - then show ?case using T by simp -qed - -lemmas mset_less_trans = mset_order.less_trans - -lemma mset_less_diff_self: "c \# B \ B - {#c#} \# B" -by (auto simp: mset_le_def mset_less_def multi_drop_mem_not_eq) - - -subsection {* Strong induction and subset induction for multisets *} - -text {* Well-foundedness of proper subset operator: *} - -text {* proper multiset subset *} -definition - mset_less_rel :: "('a multiset * 'a multiset) set" where - "mset_less_rel = {(A,B). A \# B}" - -lemma multiset_add_sub_el_shuffle: - assumes "c \# B" and "b \ c" - shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" -proof - - from `c \# B` obtain A where B: "B = A + {#c#}" - by (blast dest: multi_member_split) - have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp - then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" - by (simp add: union_ac) - then show ?thesis using B by simp -qed - -lemma wf_mset_less_rel: "wf mset_less_rel" -apply (unfold mset_less_rel_def) -apply (rule wf_measure [THEN wf_subset, where f1=size]) -apply (clarsimp simp: measure_def inv_image_def mset_less_size) -done - -text {* The induction rules: *} - -lemma full_multiset_induct [case_names less]: -assumes ih: "\B. \A. A \# B \ P A \ P B" -shows "P B" -apply (rule wf_mset_less_rel [THEN wf_induct]) -apply (rule ih, auto simp: mset_less_rel_def) -done - -lemma multi_subset_induct [consumes 2, case_names empty add]: -assumes "F \# A" - and empty: "P {#}" - and insert: "\a F. a \# A \ P F \ P (F + {#a#})" -shows "P F" -proof - - from `F \# A` - show ?thesis - proof (induct F) - show "P {#}" by fact - next - fix x F - assume P: "F \# A \ P F" and i: "F + {#x#} \# A" - show "P (F + {#x#})" - proof (rule insert) - from i show "x \# A" by (auto dest: mset_le_insertD) - from i have "F \# A" by (auto dest: mset_le_insertD) - with P show "P F" . - qed - qed -qed - -text{* A consequence: Extensionality. *} - -lemma multi_count_eq: "(\x. count A x = count B x) = (A = B)" -apply (rule iffI) - prefer 2 - apply clarsimp -apply (induct A arbitrary: B rule: full_multiset_induct) -apply (rename_tac C) -apply (case_tac B rule: multiset_cases) - apply (simp add: empty_multiset_count) -apply simp -apply (case_tac "x \# C") - apply (force dest: multi_member_split) -apply (erule_tac x = x in allE) -apply simp -done - -lemmas multi_count_ext = multi_count_eq [THEN iffD1, rule_format] - - subsection {* The fold combinator *} text {* @@ -1282,9 +1399,7 @@ lemma fold_mset_empty[simp]: "fold_mset f z {#} = z" unfolding fold_mset_def by blast -locale left_commutative = -fixes f :: "'a => 'b => 'b" -assumes left_commute: "f x (f y z) = f y (f x z)" +context fun_left_comm begin lemma fold_msetG_determ: @@ -1324,7 +1439,7 @@ have cinB: "c \# B" and binC: "b \# C" using MBb MCc diff by (auto intro: insert_noteq_member dest: sym) have "B - {#c#} \# B" using cinB by (rule mset_less_diff_self) - then have DsubM: "?D \# M" using BsubM by (blast intro: mset_less_trans) + then have DsubM: "?D \# M" using BsubM by (blast intro: mset_order.less_trans) from MBb MCc have "B + {#b#} = C + {#c#}" by blast then have [simp]: "B + {#b#} - {#c#} = C" using MBb MCc binC cinB by auto @@ -1342,7 +1457,7 @@ dest: fold_msetG.insertI [where x=b]) then have "f b d = v" using IH CsubM Cv by blast ultimately show ?thesis using x\<^isub>1 x\<^isub>2 - by (auto simp: left_commute) + by (auto simp: fun_left_comm) qed qed qed @@ -1363,7 +1478,7 @@ lemma fold_mset_insert: "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)" -apply (simp add: fold_mset_def fold_mset_insert_aux union_commute) +apply (simp add: fold_mset_def fold_mset_insert_aux add_commute) apply (rule the_equality) apply (auto cong add: conj_cong simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) @@ -1378,7 +1493,7 @@ done lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A" -by (induct A) (auto simp: fold_mset_insert left_commute [of x]) +by (induct A) (auto simp: fold_mset_insert fun_left_comm [of x]) lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z" using fold_mset_insert [of z "{#}"] by simp @@ -1389,7 +1504,7 @@ case empty then show ?case by simp next case (add A x) - have "A + {#x#} + B = (A+B) + {#x#}" by(simp add:union_ac) + have "A + {#x#} + B = (A+B) + {#x#}" by (simp add: add_ac) then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" by (simp add: fold_mset_insert) also have "\ = fold_mset f (fold_mset f z (A + {#x#})) B" @@ -1398,10 +1513,10 @@ qed lemma fold_mset_fusion: - assumes "left_commutative g" + assumes "fun_left_comm g" shows "(\x y. h (g x y) = f x (h y)) \ h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P") proof - - interpret left_commutative g by fact + interpret fun_left_comm g by (fact assms) show "PROP ?P" by (induct A) auto qed @@ -1430,11 +1545,11 @@ subsection {* Image *} -definition [code del]: - "image_mset f = fold_mset (op + o single o f) {#}" +definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where + "image_mset f = fold_mset (op + o single o f) {#}" -interpretation image_left_comm: left_commutative "op + o single o f" - proof qed (simp add:union_ac) +interpretation image_left_comm: fun_left_comm "op + o single o f" +proof qed (simp add: add_ac) lemma image_mset_empty [simp]: "image_mset f {#} = {#}" by (simp add: image_mset_def) @@ -1450,7 +1565,7 @@ "image_mset f (M+N) = image_mset f M + image_mset f N" apply (induct N) apply simp -apply (simp add: union_assoc [symmetric] image_mset_insert) +apply (simp add: add_assoc [symmetric] image_mset_insert) done lemma size_image_mset [simp]: "size (image_mset f M) = size M" @@ -1608,7 +1723,7 @@ val regroup_munion_conv = Function_Lib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus} - (map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp})) + (map (fn t => t RS eq_reflection) (@{thms add_ac} @ @{thms empty_idemp})) fun unfold_pwleq_tac i = (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st)) @@ -1629,4 +1744,31 @@ end *} -end + +subsection {* Legacy theorem bindings *} + +lemmas multi_count_eq = multiset_eq_conv_count_eq [symmetric] + +lemma union_commute: "M + N = N + (M::'a multiset)" + by (fact add_commute) + +lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" + by (fact add_assoc) + +lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" + by (fact add_left_commute) + +lemmas union_ac = union_assoc union_commute union_lcomm + +lemma union_right_cancel: "M + K = N + K \ M = (N::'a multiset)" + by (fact add_right_cancel) + +lemma union_left_cancel: "K + M = K + N \ M = (N::'a multiset)" + by (fact add_left_cancel) + +lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" + by (fact add_imp_eq) + +lemmas mset_less_trans = mset_order.less_trans + +end \ No newline at end of file diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Number_Theory/UniqueFactorization.thy --- a/src/HOL/Number_Theory/UniqueFactorization.thy Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy Fri Jan 22 15:26:29 2010 +0100 @@ -1,5 +1,4 @@ (* Title: UniqueFactorization.thy - ID: Author: Jeremy Avigad @@ -388,7 +387,7 @@ apply (subgoal_tac "multiset_prime_factorization n = Abs_multiset f") apply (unfold prime_factors_nat_def multiplicity_nat_def) - apply (simp add: set_of_def count_def Abs_multiset_inverse multiset_def) + apply (simp add: set_of_def Abs_multiset_inverse multiset_def) apply (unfold multiset_prime_factorization_def) apply (subgoal_tac "n > 0") prefer 2 @@ -401,7 +400,7 @@ apply force apply force apply force - unfolding set_of_def count_def msetprod_def + unfolding set_of_def msetprod_def apply (subgoal_tac "f : multiset") apply (auto simp only: Abs_multiset_inverse) unfolding multiset_def apply force diff -r 3b1957113753 -r dcd0fa5cc6d3 src/HOL/Tools/numeral.ML --- a/src/HOL/Tools/numeral.ML Fri Jan 22 11:42:28 2010 +0100 +++ b/src/HOL/Tools/numeral.ML Fri Jan 22 15:26:29 2010 +0100 @@ -8,7 +8,7 @@ sig val mk_cnumeral: int -> cterm val mk_cnumber: ctyp -> int -> cterm - val add_code: string -> bool -> bool -> (string -> Pretty.T) -> string -> theory -> theory + val add_code: string -> bool -> (Code_Printer.literals -> int -> string) -> string -> theory -> theory end; structure Numeral: NUMERAL = @@ -56,7 +56,7 @@ local open Basic_Code_Thingol in -fun add_code number_of negative unbounded print target thy = +fun add_code number_of negative print target thy = let fun dest_numeral pls' min' bit0' bit1' thm = let @@ -74,8 +74,7 @@ | dest_num _ = Code_Printer.eqn_error thm "Illegal numeral expression: illegal term"; in dest_num end; fun pretty literals [pls', min', bit0', bit1'] _ thm _ _ [(t, _)] = - (print o Code_Printer.literal_numeral literals unbounded - o the_default 0 o dest_numeral pls' min' bit0' bit1' thm) t; + (Code_Printer.str o print literals o the_default 0 o dest_numeral pls' min' bit0' bit1' thm) t; in thy |> Code_Target.add_syntax_const target number_of (SOME (1, ([@{const_name Int.Pls}, @{const_name Int.Min}, diff -r 3b1957113753 -r dcd0fa5cc6d3 src/Tools/Code/code_haskell.ML --- a/src/Tools/Code/code_haskell.ML Fri Jan 22 11:42:28 2010 +0100 +++ b/src/Tools/Code/code_haskell.ML Fri Jan 22 15:26:29 2010 +0100 @@ -401,11 +401,14 @@ let val s = ML_Syntax.print_char c; in if s = "'" then "\\'" else s end; + fun numeral_haskell k = if k >= 0 then string_of_int k + else Library.enclose "(" ")" (signed_string_of_int k); in Literals { literal_char = Library.enclose "'" "'" o char_haskell, literal_string = quote o translate_string char_haskell, - literal_numeral = fn unbounded => fn k => if k >= 0 then string_of_int k - else Library.enclose "(" ")" (signed_string_of_int k), + literal_numeral = numeral_haskell, + literal_positive_numeral = numeral_haskell, + literal_naive_numeral = numeral_haskell, literal_list = enum "," "[" "]", infix_cons = (5, ":") } end; diff -r 3b1957113753 -r dcd0fa5cc6d3 src/Tools/Code/code_ml.ML --- a/src/Tools/Code/code_ml.ML Fri Jan 22 11:42:28 2010 +0100 +++ b/src/Tools/Code/code_ml.ML Fri Jan 22 15:26:29 2010 +0100 @@ -354,9 +354,9 @@ val literals_sml = Literals { literal_char = prefix "#" o quote o ML_Syntax.print_char, literal_string = quote o translate_string ML_Syntax.print_char, - literal_numeral = fn unbounded => fn k => - if unbounded then "(" ^ string_of_int k ^ " : IntInf.int)" - else string_of_int k, + literal_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)", + literal_positive_numeral = fn k => "(" ^ string_of_int k ^ " : IntInf.int)", + literal_naive_numeral = string_of_int, literal_list = enum "," "[" "]", infix_cons = (7, "::") }; @@ -687,18 +687,17 @@ val s = if i < 32 orelse i = 34 orelse i = 39 orelse i = 92 orelse i > 126 then chr i else c in s end; - fun bignum_ocaml k = if k <= 1073741823 - then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")" - else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")" + fun numeral_ocaml k = if k < 0 + then "(Big_int.minus_big_int " ^ numeral_ocaml (~ k) ^ ")" + else if k <= 1073741823 + then "(Big_int.big_int_of_int " ^ string_of_int k ^ ")" + else "(Big_int.big_int_of_string " ^ quote (string_of_int k) ^ ")" in Literals { literal_char = Library.enclose "'" "'" o char_ocaml, literal_string = quote o translate_string char_ocaml, - literal_numeral = fn unbounded => fn k => if k >= 0 then - if unbounded then bignum_ocaml k - else string_of_int k - else - if unbounded then "(Big_int.minus_big_int " ^ bignum_ocaml (~ k) ^ ")" - else (Library.enclose "(" ")" o prefix "-" o string_of_int o op ~) k, + literal_numeral = numeral_ocaml, + literal_positive_numeral = numeral_ocaml, + literal_naive_numeral = numeral_ocaml, literal_list = enum ";" "[" "]", infix_cons = (6, "::") } end; @@ -975,7 +974,7 @@ ])) #> fold (Code_Target.add_reserved target_SML) ML_Syntax.reserved_names #> fold (Code_Target.add_reserved target_SML) - ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*)] + ["o" (*dictionary projections use it already*), "Fail", "div", "mod" (*standard infixes*), "IntInf"] #> fold (Code_Target.add_reserved target_OCaml) [ "and", "as", "assert", "begin", "class", "constraint", "do", "done", "downto", "else", "end", "exception", @@ -985,6 +984,6 @@ "sig", "struct", "then", "to", "true", "try", "type", "val", "virtual", "when", "while", "with" ] - #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod"]; + #> fold (Code_Target.add_reserved target_OCaml) ["failwith", "mod", "Big_int"]; end; (*struct*) diff -r 3b1957113753 -r dcd0fa5cc6d3 src/Tools/Code/code_printer.ML --- a/src/Tools/Code/code_printer.ML Fri Jan 22 11:42:28 2010 +0100 +++ b/src/Tools/Code/code_printer.ML Fri Jan 22 15:26:29 2010 +0100 @@ -39,12 +39,15 @@ type literals val Literals: { literal_char: string -> string, literal_string: string -> string, - literal_numeral: bool -> int -> string, + literal_numeral: int -> string, literal_positive_numeral: int -> string, + literal_naive_numeral: int -> string, literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string } -> literals val literal_char: literals -> string -> string val literal_string: literals -> string -> string - val literal_numeral: literals -> bool -> int -> string + val literal_numeral: literals -> int -> string + val literal_positive_numeral: literals -> int -> string + val literal_naive_numeral: literals -> int -> string val literal_list: literals -> Pretty.T list -> Pretty.T val infix_cons: literals -> int * string @@ -163,7 +166,9 @@ datatype literals = Literals of { literal_char: string -> string, literal_string: string -> string, - literal_numeral: bool -> int -> string, + literal_numeral: int -> string, + literal_positive_numeral: int -> string, + literal_naive_numeral: int -> string, literal_list: Pretty.T list -> Pretty.T, infix_cons: int * string }; @@ -173,6 +178,8 @@ val literal_char = #literal_char o dest_Literals; val literal_string = #literal_string o dest_Literals; val literal_numeral = #literal_numeral o dest_Literals; +val literal_positive_numeral = #literal_positive_numeral o dest_Literals; +val literal_naive_numeral = #literal_naive_numeral o dest_Literals; val literal_list = #literal_list o dest_Literals; val infix_cons = #infix_cons o dest_Literals; diff -r 3b1957113753 -r dcd0fa5cc6d3 src/Tools/Code/code_scala.ML --- a/src/Tools/Code/code_scala.ML Fri Jan 22 11:42:28 2010 +0100 +++ b/src/Tools/Code/code_scala.ML Fri Jan 22 15:26:29 2010 +0100 @@ -404,17 +404,18 @@ let val s = ML_Syntax.print_char c; in if s = "'" then "\\'" else s end; - fun bigint_scala k = "(" ^ (if k <= 2147483647 - then string_of_int k else quote (string_of_int k)) ^ ")" + fun numeral_scala k = if k < 0 + then if k <= 2147483647 then "- " ^ string_of_int (~ k) + else quote ("- " ^ string_of_int (~ k)) + else if k <= 2147483647 then string_of_int k + else quote (string_of_int k) in Literals { literal_char = Library.enclose "'" "'" o char_scala, literal_string = quote o translate_string char_scala, - literal_numeral = fn unbounded => fn k => if k >= 0 then - if unbounded then bigint_scala k - else Library.enclose "(" ")" (string_of_int k) - else - if unbounded then "(- " ^ bigint_scala (~ k) ^ ")" - else Library.enclose "(" ")" (signed_string_of_int k), + literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")", + literal_positive_numeral = fn k => "Nat.Nat(" ^ numeral_scala k ^ ")", + literal_naive_numeral = fn k => if k >= 0 + then string_of_int k else "(- " ^ string_of_int (~ k) ^ ")", literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps], infix_cons = (6, "::") } end; @@ -442,7 +443,7 @@ "true", "type", "val", "var", "while", "with" ] #> fold (Code_Target.add_reserved target) [ - "error", "apply", "List", "Nil" + "error", "apply", "List", "Nil", "BigInt" ]; end; (*struct*)