# HG changeset patch # User wenzelm # Date 1470859536 -7200 # Node ID f90e3926e627a957ce98a2a1bdd99053c5c15e34 # Parent 4453cfb745e5b4c66d0c7ab2e26fcb0871f7aef9 misc tuning and modernization; diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Fun_Def.thy --- a/src/HOL/Fun_Def.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Fun_Def.thy Wed Aug 10 22:05:36 2016 +0200 @@ -5,32 +5,29 @@ section \Function Definitions and Termination Proofs\ theory Fun_Def -imports Basic_BNF_LFPs Partial_Function SAT -keywords - "function" "termination" :: thy_goal and - "fun" "fun_cases" :: thy_decl + imports Basic_BNF_LFPs Partial_Function SAT + keywords + "function" "termination" :: thy_goal and + "fun" "fun_cases" :: thy_decl begin subsection \Definitions with default value\ -definition - THE_default :: "'a \ ('a \ bool) \ 'a" where - "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" +definition THE_default :: "'a \ ('a \ bool) \ 'a" + where "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" lemma THE_defaultI': "\!x. P x \ P (THE_default d P)" by (simp add: theI' THE_default_def) -lemma THE_default1_equality: - "\\!x. P x; P a\ \ THE_default d P = a" +lemma THE_default1_equality: "\!x. P x \ P a \ THE_default d P = a" by (simp add: the1_equality THE_default_def) -lemma THE_default_none: - "\(\!x. P x) \ THE_default d P = d" - by (simp add:THE_default_def) +lemma THE_default_none: "\ (\!x. P x) \ THE_default d P = d" + by (simp add: THE_default_def) lemma fundef_ex1_existence: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "G x (f x)" apply (simp only: f_def) @@ -39,7 +36,7 @@ done lemma fundef_ex1_uniqueness: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" assumes elm: "G x (h x)" shows "h x = f x" @@ -50,7 +47,7 @@ done lemma fundef_ex1_iff: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "(G x y) = (f x = y)" apply (auto simp:ex1 f_def THE_default1_equality) @@ -59,7 +56,7 @@ done lemma fundef_default_value: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes graph: "\x y. G x y \ D x" assumes "\ D x" shows "f x = d x" @@ -67,21 +64,17 @@ have "\(\y. G x y)" proof assume "\y. G x y" - hence "D x" using graph .. + then have "D x" using graph .. with \\ D x\ show False .. qed - hence "\(\!y. G x y)" by blast - - thus ?thesis - unfolding f_def - by (rule THE_default_none) + then have "\(\!y. G x y)" by blast + then show ?thesis + unfolding f_def by (rule THE_default_none) qed -definition in_rel_def[simp]: - "in_rel R x y == (x, y) \ R" +definition in_rel_def[simp]: "in_rel R x y \ (x, y) \ R" -lemma wf_in_rel: - "wf R \ wfP (in_rel R)" +lemma wf_in_rel: "wf R \ wfP (in_rel R)" by (simp add: wfP_def) ML_file "Tools/Function/function_core.ML" @@ -112,18 +105,19 @@ subsection \Measure functions\ inductive is_measure :: "('a \ nat) \ bool" -where is_measure_trivial: "is_measure f" + where is_measure_trivial: "is_measure f" named_theorems measure_function "rules that guide the heuristic generation of measure functions" ML_file "Tools/Function/measure_functions.ML" lemma measure_size[measure_function]: "is_measure size" -by (rule is_measure_trivial) + by (rule is_measure_trivial) lemma measure_fst[measure_function]: "is_measure f \ is_measure (\p. f (fst p))" -by (rule is_measure_trivial) + by (rule is_measure_trivial) + lemma measure_snd[measure_function]: "is_measure f \ is_measure (\p. f (snd p))" -by (rule is_measure_trivial) + by (rule is_measure_trivial) ML_file "Tools/Function/lexicographic_order.ML" @@ -135,8 +129,7 @@ subsection \Congruence rules\ -lemma let_cong [fundef_cong]: - "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" +lemma let_cong [fundef_cong]: "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" unfolding Let_def by blast lemmas [fundef_cong] = @@ -144,13 +137,11 @@ bex_cong ball_cong imp_cong map_option_cong Option.bind_cong lemma split_cong [fundef_cong]: - "(\x y. (x, y) = q \ f x y = g x y) \ p = q - \ case_prod f p = case_prod g q" + "(\x y. (x, y) = q \ f x y = g x y) \ p = q \ case_prod f p = case_prod g q" by (auto simp: split_def) -lemma comp_cong [fundef_cong]: - "f (g x) = f' (g' x') \ (f o g) x = (f' o g') x'" - unfolding o_apply . +lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \ (f \ g) x = (f' \ g') x'" + by (simp only: o_apply) subsection \Simp rules for termination proofs\ @@ -163,31 +154,25 @@ less_imp_le_nat[termination_simp] le_imp_less_Suc[termination_simp] -lemma size_prod_simp[termination_simp]: - "size_prod f g p = f (fst p) + g (snd p) + Suc 0" -by (induct p) auto +lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0" + by (induct p) auto subsection \Decomposition\ -lemma less_by_empty: - "A = {} \ A \ B" -and union_comp_emptyL: - "\ A O C = {}; B O C = {} \ \ (A \ B) O C = {}" -and union_comp_emptyR: - "\ A O B = {}; A O C = {} \ \ A O (B \ C) = {}" -and wf_no_loop: - "R O R = {} \ wf R" -by (auto simp add: wf_comp_self[of R]) +lemma less_by_empty: "A = {} \ A \ B" + and union_comp_emptyL: "A O C = {} \ B O C = {} \ (A \ B) O C = {}" + and union_comp_emptyR: "A O B = {} \ A O C = {} \ A O (B \ C) = {}" + and wf_no_loop: "R O R = {} \ wf R" + by (auto simp add: wf_comp_self [of R]) subsection \Reduction pairs\ -definition - "reduction_pair P = (wf (fst P) \ fst P O snd P \ fst P)" +definition "reduction_pair P \ wf (fst P) \ fst P O snd P \ fst P" lemma reduction_pairI[intro]: "wf R \ R O S \ R \ reduction_pair (R, S)" -unfolding reduction_pair_def by auto + by (auto simp: reduction_pair_def) lemma reduction_pair_lemma: assumes rp: "reduction_pair P" @@ -204,13 +189,10 @@ ultimately show ?thesis by (rule wf_subset) qed -definition - "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" +definition "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" -lemma rp_inv_image_rp: - "reduction_pair P \ reduction_pair (rp_inv_image P f)" - unfolding reduction_pair_def rp_inv_image_def split_def - by force +lemma rp_inv_image_rp: "reduction_pair P \ reduction_pair (rp_inv_image P f)" + unfolding reduction_pair_def rp_inv_image_def split_def by force subsection \Concrete orders for SCNP termination proofs\ @@ -230,70 +212,70 @@ and pair_leqI2: "a \ b \ s \ t \ ((a, s), (b, t)) \ pair_leq" and pair_lessI1: "a < b \ ((a, s), (b, t)) \ pair_less" and pair_lessI2: "a \ b \ s < t \ ((a, s), (b, t)) \ pair_less" - unfolding pair_leq_def pair_less_def by auto + by (auto simp: pair_leq_def pair_less_def) text \Introduction rules for max\ -lemma smax_emptyI: - "finite Y \ Y \ {} \ ({}, Y) \ max_strict" +lemma smax_emptyI: "finite Y \ Y \ {} \ ({}, Y) \ max_strict" and smax_insertI: - "\y \ Y; (x, y) \ pair_less; (X, Y) \ max_strict\ \ (insert x X, Y) \ max_strict" - and wmax_emptyI: - "finite X \ ({}, X) \ max_weak" + "y \ Y \ (x, y) \ pair_less \ (X, Y) \ max_strict \ (insert x X, Y) \ max_strict" + and wmax_emptyI: "finite X \ ({}, X) \ max_weak" and wmax_insertI: - "\y \ YS; (x, y) \ pair_leq; (XS, YS) \ max_weak\ \ (insert x XS, YS) \ max_weak" -unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases) + "y \ YS \ (x, y) \ pair_leq \ (XS, YS) \ max_weak \ (insert x XS, YS) \ max_weak" + by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases) text \Introduction rules for min\ -lemma smin_emptyI: - "X \ {} \ (X, {}) \ min_strict" +lemma smin_emptyI: "X \ {} \ (X, {}) \ min_strict" and smin_insertI: - "\x \ XS; (x, y) \ pair_less; (XS, YS) \ min_strict\ \ (XS, insert y YS) \ min_strict" - and wmin_emptyI: - "(X, {}) \ min_weak" + "x \ XS \ (x, y) \ pair_less \ (XS, YS) \ min_strict \ (XS, insert y YS) \ min_strict" + and wmin_emptyI: "(X, {}) \ min_weak" and wmin_insertI: - "\x \ XS; (x, y) \ pair_leq; (XS, YS) \ min_weak\ \ (XS, insert y YS) \ min_weak" -by (auto simp: min_strict_def min_weak_def min_ext_def) + "x \ XS \ (x, y) \ pair_leq \ (XS, YS) \ min_weak \ (XS, insert y YS) \ min_weak" + by (auto simp: min_strict_def min_weak_def min_ext_def) -text \Reduction Pairs\ +text \Reduction Pairs.\ lemma max_ext_compat: assumes "R O S \ R" - shows "max_ext R O (max_ext S \ {({},{})}) \ max_ext R" -using assms -apply auto -apply (elim max_ext.cases) -apply rule -apply auto[3] -apply (drule_tac x=xa in meta_spec) -apply simp -apply (erule bexE) -apply (drule_tac x=xb in meta_spec) -by auto + shows "max_ext R O (max_ext S \ {({}, {})}) \ max_ext R" + using assms + apply auto + apply (elim max_ext.cases) + apply rule + apply auto[3] + apply (drule_tac x=xa in meta_spec) + apply simp + apply (erule bexE) + apply (drule_tac x=xb in meta_spec) + apply auto + done lemma max_rpair_set: "reduction_pair (max_strict, max_weak)" unfolding max_strict_def max_weak_def -apply (intro reduction_pairI max_ext_wf) -apply simp -apply (rule max_ext_compat) -by (auto simp: pair_less_def pair_leq_def) + apply (intro reduction_pairI max_ext_wf) + apply simp + apply (rule max_ext_compat) + apply (auto simp: pair_less_def pair_leq_def) + done lemma min_ext_compat: assumes "R O S \ R" shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" -using assms -apply (auto simp: min_ext_def) -apply (drule_tac x=ya in bspec, assumption) -apply (erule bexE) -apply (drule_tac x=xc in bspec) -apply assumption -by auto + using assms + apply (auto simp: min_ext_def) + apply (drule_tac x=ya in bspec, assumption) + apply (erule bexE) + apply (drule_tac x=xc in bspec) + apply assumption + apply auto + done lemma min_rpair_set: "reduction_pair (min_strict, min_weak)" unfolding min_strict_def min_weak_def -apply (intro reduction_pairI min_ext_wf) -apply simp -apply (rule min_ext_compat) -by (auto simp: pair_less_def pair_leq_def) + apply (intro reduction_pairI min_ext_wf) + apply simp + apply (rule min_ext_compat) + apply (auto simp: pair_less_def pair_leq_def) + done subsection \Tool setup\ diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Groups_Big.thy Wed Aug 10 22:05:36 2016 +0200 @@ -1,12 +1,14 @@ (* Title: HOL/Groups_Big.thy - Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel - with contributions by Jeremy Avigad + Author: Tobias Nipkow + Author: Lawrence C Paulson + Author: Markus Wenzel + Author: Jeremy Avigad *) section \Big sum and product over finite (non-empty) sets\ theory Groups_Big -imports Finite_Set Power + imports Finite_Set Power begin subsection \Generic monoid operation over a set\ @@ -21,60 +23,53 @@ by (fact comp_comp_fun_commute) definition F :: "('b \ 'a) \ 'b set \ 'a" -where - eq_fold: "F g A = Finite_Set.fold (f \ g) \<^bold>1 A" + where eq_fold: "F g A = Finite_Set.fold (f \ g) \<^bold>1 A" -lemma infinite [simp]: - "\ finite A \ F g A = \<^bold>1" +lemma infinite [simp]: "\ finite A \ F g A = \<^bold>1" by (simp add: eq_fold) -lemma empty [simp]: - "F g {} = \<^bold>1" +lemma empty [simp]: "F g {} = \<^bold>1" by (simp add: eq_fold) -lemma insert [simp]: - assumes "finite A" and "x \ A" - shows "F g (insert x A) = g x \<^bold>* F g A" - using assms by (simp add: eq_fold) +lemma insert [simp]: "finite A \ x \ A \ F g (insert x A) = g x \<^bold>* F g A" + by (simp add: eq_fold) lemma remove: assumes "finite A" and "x \ A" shows "F g A = g x \<^bold>* F g (A - {x})" proof - - from \x \ A\ obtain B where A: "A = insert x B" and "x \ B" + from \x \ A\ obtain B where B: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) - moreover from \finite A\ A have "finite B" by simp + moreover from \finite A\ B have "finite B" by simp ultimately show ?thesis by simp qed -lemma insert_remove: - assumes "finite A" - shows "F g (insert x A) = g x \<^bold>* F g (A - {x})" - using assms by (cases "x \ A") (simp_all add: remove insert_absorb) +lemma insert_remove: "finite A \ F g (insert x A) = g x \<^bold>* F g (A - {x})" + by (cases "x \ A") (simp_all add: remove insert_absorb) -lemma neutral: - assumes "\x\A. g x = \<^bold>1" - shows "F g A = \<^bold>1" - using assms by (induct A rule: infinite_finite_induct) simp_all +lemma neutral: "\x\A. g x = \<^bold>1 \ F g A = \<^bold>1" + by (induct A rule: infinite_finite_induct) simp_all -lemma neutral_const [simp]: - "F (\_. \<^bold>1) A = \<^bold>1" +lemma neutral_const [simp]: "F (\_. \<^bold>1) A = \<^bold>1" by (simp add: neutral) lemma union_inter: assumes "finite A" and "finite B" shows "F g (A \ B) \<^bold>* F g (A \ B) = F g A \<^bold>* F g B" \ \The reversed orientation looks more natural, but LOOPS as a simprule!\ -using assms proof (induct A) - case empty then show ?case by simp + using assms +proof (induct A) + case empty + then show ?case by simp next - case (insert x A) then show ?case - by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) + case (insert x A) + then show ?case + by (auto simp: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) qed corollary union_inter_neutral: assumes "finite A" and "finite B" - and I0: "\x \ A \ B. g x = \<^bold>1" + and "\x \ A \ B. g x = \<^bold>1" shows "F g (A \ B) = F g A \<^bold>* F g B" using assms by (simp add: union_inter [symmetric] neutral) @@ -90,7 +85,8 @@ proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto - with assms show ?thesis by simp (subst union_disjoint, auto)+ + with assms show ?thesis + by simp (subst union_disjoint, auto)+ qed lemma subset_diff: @@ -116,9 +112,15 @@ proof - from assms have "\a\A. g a \ \<^bold>1" proof (induct A rule: infinite_finite_induct) + case infinite + then show ?case by simp + next + case empty + then show ?case by simp + next case (insert a A) - then show ?case by simp (rule, simp) - qed simp_all + then show ?case by fastforce + qed with that show thesis by blast qed @@ -127,9 +129,11 @@ shows "F g (h ` A) = F (g \ h) A" proof (cases "finite A") case True - with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) + with assms show ?thesis + by (simp add: eq_fold fold_image comp_assoc) next - case False with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) + case False + with assms have "\ finite (h ` A)" by (blast dest: finite_imageD) with False show ?thesis by simp qed @@ -143,7 +147,7 @@ lemma strong_cong [cong]: assumes "A = B" "\x. x \ B =simp=> g x = h x" shows "F (\x. g x) A = F (\x. h x) B" - by (rule cong) (insert assms, simp_all add: simp_implies_def) + by (rule cong) (use assms in \simp_all add: simp_implies_def\) lemma reindex_cong: assumes "inj_on l B" @@ -154,55 +158,64 @@ lemma UNION_disjoint: assumes "finite I" and "\i\I. finite (A i)" - and "\i\I. \j\I. i \ j \ A i \ A j = {}" + and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "F g (UNION I A) = F (\x. F g (A x)) I" -apply (insert assms) -apply (induct rule: finite_induct) -apply simp -apply atomize -apply (subgoal_tac "\i\Fa. x \ i") - prefer 2 apply blast -apply (subgoal_tac "A x Int UNION Fa A = {}") - prefer 2 apply blast -apply (simp add: union_disjoint) -done + apply (insert assms) + apply (induct rule: finite_induct) + apply simp + apply atomize + apply (subgoal_tac "\i\Fa. x \ i") + prefer 2 apply blast + apply (subgoal_tac "A x \ UNION Fa A = {}") + prefer 2 apply blast + apply (simp add: union_disjoint) + done lemma Union_disjoint: assumes "\A\C. finite A" "\A\C. \B\C. A \ B \ A \ B = {}" shows "F g (\C) = (F \ F) g C" -proof cases - assume "finite C" - from UNION_disjoint [OF this assms] - show ?thesis by simp -qed (auto dest: finite_UnionD intro: infinite) +proof (cases "finite C") + case True + from UNION_disjoint [OF this assms] show ?thesis by simp +next + case False + then show ?thesis by (auto dest: finite_UnionD intro: infinite) +qed -lemma distrib: - "F (\x. g x \<^bold>* h x) A = F g A \<^bold>* F h A" +lemma distrib: "F (\x. g x \<^bold>* h x) A = F g A \<^bold>* F h A" by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) lemma Sigma: "finite A \ \x\A. finite (B x) \ F (\x. F (g x) (B x)) A = F (case_prod g) (SIGMA x:A. B x)" -apply (subst Sigma_def) -apply (subst UNION_disjoint, assumption, simp) - apply blast -apply (rule cong) -apply rule -apply (simp add: fun_eq_iff) -apply (subst UNION_disjoint, simp, simp) - apply blast -apply (simp add: comp_def) -done + apply (subst Sigma_def) + apply (subst UNION_disjoint) + apply assumption + apply simp + apply blast + apply (rule cong) + apply rule + apply (simp add: fun_eq_iff) + apply (subst UNION_disjoint) + apply simp + apply simp + apply blast + apply (simp add: comp_def) + done lemma related: assumes Re: "R \<^bold>1 \<^bold>1" - and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 \<^bold>* y1) (x2 \<^bold>* y2)" - and fS: "finite S" and Rfg: "\x\S. R (h x) (g x)" + and Rop: "\x1 y1 x2 y2. R x1 x2 \ R y1 y2 \ R (x1 \<^bold>* y1) (x2 \<^bold>* y2)" + and fin: "finite S" + and R_h_g: "\x\S. R (h x) (g x)" shows "R (F h S) (F g S)" - using fS by (rule finite_subset_induct) (insert assms, auto) + using fin by (rule finite_subset_induct) (use assms in auto) lemma mono_neutral_cong_left: - assumes "finite T" and "S \ T" and "\i \ T - S. h i = \<^bold>1" - and "\x. x \ S \ g x = h x" shows "F g S = F h T" + assumes "finite T" + and "S \ T" + and "\i \ T - S. h i = \<^bold>1" + and "\x. x \ S \ g x = h x" + shows "F g S = F h T" proof- have eq: "T = S \ (T - S)" using \S \ T\ by blast have d: "S \ (T - S) = {}" using \S \ T\ by blast @@ -213,16 +226,14 @@ qed lemma mono_neutral_cong_right: - "\ finite T; S \ T; \i \ T - S. g i = \<^bold>1; \x. x \ S \ g x = h x \ - \ F g T = F h S" + "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ (\x. x \ S \ g x = h x) \ + F g T = F h S" by (auto intro!: mono_neutral_cong_left [symmetric]) -lemma mono_neutral_left: - "\ finite T; S \ T; \i \ T - S. g i = \<^bold>1 \ \ F g S = F g T" +lemma mono_neutral_left: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g S = F g T" by (blast intro: mono_neutral_cong_left) -lemma mono_neutral_right: - "\ finite T; S \ T; \i \ T - S. g i = \<^bold>1 \ \ F g T = F g S" +lemma mono_neutral_right: "finite T \ S \ T \ \i \ T - S. g i = \<^bold>1 \ F g T = F g S" by (blast intro!: mono_neutral_left [symmetric]) lemma reindex_bij_betw: "bij_betw h S T \ F (\x. g (h x)) S = F g T" @@ -256,10 +267,9 @@ proof - have [simp]: "finite S \ finite T" using bij_betw_finite[OF bij] fin by auto - show ?thesis - proof cases - assume "finite S" + proof (cases "finite S") + case True with nn have "F (\x. g (h x)) S = F (\x. g (h x)) (S - S')" by (intro mono_neutral_cong_right) auto also have "\ = F g (T - T')" @@ -267,17 +277,20 @@ also have "\ = F g T" using nn \finite S\ by (intro mono_neutral_cong_left) auto finally show ?thesis . - qed simp + next + case False + then show ?thesis by simp + qed qed lemma reindex_nontrivial: assumes "finite A" - and nz: "\x y. x \ A \ y \ A \ x \ y \ h x = h y \ g (h x) = \<^bold>1" + and nz: "\x y. x \ A \ y \ A \ x \ y \ h x = h y \ g (h x) = \<^bold>1" shows "F g (h ` A) = F (g \ h) A" proof (subst reindex_bij_betw_not_neutral [symmetric]) show "bij_betw h (A - {x \ A. (g \ h) x = \<^bold>1}) (h ` A - h ` {x \ A. (g \ h) x = \<^bold>1})" using nz by (auto intro!: inj_onI simp: bij_betw_def) -qed (insert \finite A\, auto) +qed (use \finite A\ in auto) lemma reindex_bij_witness_not_neutral: assumes fin: "finite S'" "finite T'" @@ -305,69 +318,66 @@ lemma delta: assumes fS: "finite S" shows "F (\k. if k = a then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" -proof- - let ?f = "(\k. if k=a then b k else \<^bold>1)" - { assume a: "a \ S" - hence "\k\S. ?f k = \<^bold>1" by simp - hence ?thesis using a by simp } - moreover - { assume a: "a \ S" +proof - + let ?f = "(\k. if k = a then b k else \<^bold>1)" + show ?thesis + proof (cases "a \ S") + case False + then have "\k\S. ?f k = \<^bold>1" by simp + with False show ?thesis by simp + next + case True let ?A = "S - {a}" let ?B = "{a}" - have eq: "S = ?A \ ?B" using a by blast + from True have eq: "S = ?A \ ?B" by blast have dj: "?A \ ?B = {}" by simp from fS have fAB: "finite ?A" "finite ?B" by auto have "F ?f S = F ?f ?A \<^bold>* F ?f ?B" - using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] - by simp - then have ?thesis using a by simp } - ultimately show ?thesis by blast + using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] by simp + with True show ?thesis by simp + qed qed lemma delta': - assumes fS: "finite S" + assumes fin: "finite S" shows "F (\k. if a = k then b k else \<^bold>1) S = (if a \ S then b a else \<^bold>1)" - using delta [OF fS, of a b, symmetric] by (auto intro: cong) + using delta [OF fin, of a b, symmetric] by (auto intro: cong) lemma If_cases: fixes P :: "'b \ bool" and g h :: "'b \ 'a" - assumes fA: "finite A" - shows "F (\x. if P x then h x else g x) A = - F h (A \ {x. P x}) \<^bold>* F g (A \ - {x. P x})" + assumes fin: "finite A" + shows "F (\x. if P x then h x else g x) A = F h (A \ {x. P x}) \<^bold>* F g (A \ - {x. P x})" proof - - have a: "A = A \ {x. P x} \ A \ -{x. P x}" - "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" + have a: "A = A \ {x. P x} \ A \ -{x. P x}" "(A \ {x. P x}) \ (A \ -{x. P x}) = {}" by blast+ - from fA - have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto + from fin have f: "finite (A \ {x. P x})" "finite (A \ -{x. P x})" by auto let ?g = "\x. if P x then h x else g x" - from union_disjoint [OF f a(2), of ?g] a(1) - show ?thesis + from union_disjoint [OF f a(2), of ?g] a(1) show ?thesis by (subst (1 2) cong) simp_all qed -lemma cartesian_product: - "F (\x. F (g x) B) A = F (case_prod g) (A \ B)" -apply (rule sym) -apply (cases "finite A") - apply (cases "finite B") - apply (simp add: Sigma) - apply (cases "A={}", simp) - apply simp -apply (auto intro: infinite dest: finite_cartesian_productD2) -apply (cases "B = {}") apply (auto intro: infinite dest: finite_cartesian_productD1) -done +lemma cartesian_product: "F (\x. F (g x) B) A = F (case_prod g) (A \ B)" + apply (rule sym) + apply (cases "finite A") + apply (cases "finite B") + apply (simp add: Sigma) + apply (cases "A = {}") + apply simp + apply simp + apply (auto intro: infinite dest: finite_cartesian_productD2) + apply (cases "B = {}") + apply (auto intro: infinite dest: finite_cartesian_productD1) + done lemma inter_restrict: assumes "finite A" shows "F g (A \ B) = F (\x. if x \ B then g x else \<^bold>1) A" proof - let ?g = "\x. if x \ A \ B then g x else \<^bold>1" - have "\i\A - A \ B. (if i \ A \ B then g i else \<^bold>1) = \<^bold>1" - by simp + have "\i\A - A \ B. (if i \ A \ B then g i else \<^bold>1) = \<^bold>1" by simp moreover have "A \ B \ A" by blast - ultimately have "F ?g (A \ B) = F ?g A" using \finite A\ - by (intro mono_neutral_left) auto + ultimately have "F ?g (A \ B) = F ?g A" + using \finite A\ by (intro mono_neutral_left) auto then show ?thesis by simp qed @@ -377,27 +387,28 @@ lemma Union_comp: assumes "\A \ B. finite A" - and "\A1 A2 x. A1 \ B \ A2 \ B \ A1 \ A2 \ x \ A1 \ x \ A2 \ g x = \<^bold>1" + and "\A1 A2 x. A1 \ B \ A2 \ B \ A1 \ A2 \ x \ A1 \ x \ A2 \ g x = \<^bold>1" shows "F g (\B) = (F \ F) g B" -using assms proof (induct B rule: infinite_finite_induct) + using assms +proof (induct B rule: infinite_finite_induct) case (infinite A) then have "\ finite (\A)" by (blast dest: finite_UnionD) with infinite show ?case by simp next - case empty then show ?case by simp + case empty + then show ?case by simp next case (insert A B) then have "finite A" "finite B" "finite (\B)" "A \ B" and "\x\A \ \B. g x = \<^bold>1" - and H: "F g (\B) = (F o F) g B" by auto + and H: "F g (\B) = (F \ F) g B" by auto then have "F g (A \ \B) = F g A \<^bold>* F g (\B)" by (simp add: union_inter_neutral) with \finite B\ \A \ B\ show ?case by (simp add: H) qed -lemma commute: - "F (\i. F (g i) B) A = F (\j. F (\i. g i j) A) B" +lemma commute: "F (\i. F (g i) B) A = F (\j. F (\i. g i j) A) B" unfolding cartesian_product by (rule reindex_bij_witness [where i = "\(i, j). (j, i)" and j = "\(i, j). (j, i)"]) auto @@ -412,13 +423,11 @@ shows "F g (A <+> B) = F (g \ Inl) A \<^bold>* F (g \ Inr) B" proof - have "A <+> B = Inl ` A \ Inr ` B" by auto - moreover from fin have "finite (Inl ` A :: ('b + 'c) set)" "finite (Inr ` B :: ('b + 'c) set)" - by auto - moreover have "Inl ` A \ Inr ` B = ({} :: ('b + 'c) set)" by auto - moreover have "inj_on (Inl :: 'b \ 'b + 'c) A" "inj_on (Inr :: 'c \ 'b + 'c) B" - by (auto intro: inj_onI) - ultimately show ?thesis using fin - by (simp add: union_disjoint reindex) + moreover from fin have "finite (Inl ` A)" "finite (Inr ` B)" by auto + moreover have "Inl ` A \ Inr ` B = {}" by auto + moreover have "inj_on Inl A" "inj_on Inr B" by (auto intro: inj_onI) + ultimately show ?thesis + using fin by (simp add: union_disjoint reindex) qed lemma same_carrier: @@ -427,22 +436,22 @@ assumes trivial: "\a. a \ C - A \ g a = \<^bold>1" "\b. b \ C - B \ h b = \<^bold>1" shows "F g A = F h B \ F g C = F h C" proof - - from \finite C\ subset have - "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)" - by (auto elim: finite_subset) + have "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)" + using \finite C\ subset by (auto elim: finite_subset) from subset have [simp]: "A - (C - A) = A" by auto from subset have [simp]: "B - (C - B) = B" by auto from subset have "C = A \ (C - A)" by auto then have "F g C = F g (A \ (C - A))" by simp also have "\ = F g (A - (C - A)) \<^bold>* F g (C - A - A) \<^bold>* F g (A \ (C - A))" using \finite A\ \finite (C - A)\ by (simp only: union_diff2) - finally have P: "F g C = F g A" using trivial by simp + finally have *: "F g C = F g A" using trivial by simp from subset have "C = B \ (C - B)" by auto then have "F h C = F h (B \ (C - B))" by simp also have "\ = F h (B - (C - B)) \<^bold>* F h (C - B - B) \<^bold>* F h (B \ (C - B))" using \finite B\ \finite (C - B)\ by (simp only: union_diff2) - finally have Q: "F h C = F h B" using trivial by simp - from P Q show ?thesis by simp + finally have "F h C = F h B" + using trivial by simp + with * show ?thesis by simp qed lemma same_carrierI: @@ -462,8 +471,7 @@ begin sublocale setsum: comm_monoid_set plus 0 -defines - setsum = setsum.F .. + defines setsum = setsum.F .. abbreviation Setsum ("\_" [1000] 999) where "\A \ setsum (\x. x) A" @@ -504,27 +512,28 @@ in [(@{const_syntax setsum}, K setsum_tr')] end \ -text \TODO generalization candidates\ +(* TODO generalization candidates *) lemma (in comm_monoid_add) setsum_image_gen: - assumes fS: "finite S" + assumes fin: "finite S" shows "setsum g S = setsum (\y. setsum g {x. x \ S \ f x = y}) (f ` S)" -proof- - { fix x assume "x \ S" then have "{y. y\ f`S \ f x = y} = {f x}" by auto } - hence "setsum g S = setsum (\x. setsum (\y. g x) {y. y\ f`S \ f x = y}) S" +proof - + have "{y. y\ f`S \ f x = y} = {f x}" if "x \ S" for x + using that by auto + then have "setsum g S = setsum (\x. setsum (\y. g x) {y. y\ f`S \ f x = y}) S" by simp also have "\ = setsum (\y. setsum g {x. x \ S \ f x = y}) (f ` S)" - by (rule setsum.commute_restrict [OF fS finite_imageI [OF fS]]) + by (rule setsum.commute_restrict [OF fin finite_imageI [OF fin]]) finally show ?thesis . qed subsubsection \Properties in more restricted classes of structures\ -lemma setsum_Un: "finite A ==> finite B ==> - (setsum f (A Un B) :: 'a :: ab_group_add) = - setsum f A + setsum f B - setsum f (A Int B)" -by (subst setsum.union_inter [symmetric], auto simp add: algebra_simps) +lemma setsum_Un: + "finite A \ finite B \ setsum f (A \ B) = setsum f A + setsum f B - setsum f (A \ B)" + for f :: "'b \ 'a::ab_group_add" + by (subst setsum.union_inter [symmetric]) (auto simp add: algebra_simps) lemma setsum_Un2: assumes "finite (A \ B)" @@ -532,26 +541,30 @@ proof - have "A \ B = A - B \ (B - A) \ A \ B" by auto - with assms show ?thesis by simp (subst setsum.union_disjoint, auto)+ + with assms show ?thesis + by simp (subst setsum.union_disjoint, auto)+ qed -lemma setsum_diff1: "finite A \ - (setsum f (A - {a}) :: ('a::ab_group_add)) = - (if a:A then setsum f A - f a else setsum f A)" -by (erule finite_induct) (auto simp add: insert_Diff_if) +lemma setsum_diff1: + fixes f :: "'b \ 'a::ab_group_add" + assumes "finite A" + shows "setsum f (A - {a}) = (if a \ A then setsum f A - f a else setsum f A)" + using assms by induct (auto simp: insert_Diff_if) lemma setsum_diff: - assumes le: "finite A" "B \ A" - shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))" + fixes f :: "'b \ 'a::ab_group_add" + assumes "finite A" "B \ A" + shows "setsum f (A - B) = setsum f A - setsum f B" proof - - from le have finiteB: "finite B" using finite_subset by auto - show ?thesis using finiteB le + from assms(2,1) have "finite B" by (rule finite_subset) + from this \B \ A\ + show ?thesis proof induct case empty - thus ?case by auto + thus ?case by simp next case (insert x F) - thus ?case using le finiteB + with \finite A\ \finite B\ show ?case by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb) qed qed @@ -561,45 +574,52 @@ shows "(\i\K. f i) \ (\i\K. g i)" proof (cases "finite K") case True - thus ?thesis using le + from this le show ?thesis proof induct case empty - thus ?case by simp + then show ?case by simp next case insert - thus ?case using add_mono by fastforce + then show ?case using add_mono by fastforce qed next - case False then show ?thesis by simp + case False + then show ?thesis by simp qed lemma (in strict_ordered_comm_monoid_add) setsum_strict_mono: - assumes "finite A" "A \ {}" and "\x. x \ A \ f x < g x" + assumes "finite A" "A \ {}" + and "\x. x \ A \ f x < g x" shows "setsum f A < setsum g A" using assms proof (induct rule: finite_ne_induct) - case singleton thus ?case by simp + case singleton + then show ?case by simp next - case insert thus ?case by (auto simp: add_strict_mono) + case insert + then show ?case by (auto simp: add_strict_mono) qed lemma setsum_strict_mono_ex1: fixes f g :: "'i \ 'a::ordered_cancel_comm_monoid_add" - assumes "finite A" and "\x\A. f x \ g x" and "\a\A. f a < g a" + assumes "finite A" + and "\x\A. f x \ g x" + and "\a\A. f a < g a" shows "setsum f A < setsum g A" proof- - from assms(3) obtain a where a: "a:A" "f a < g a" by blast - have "setsum f A = setsum f ((A-{a}) \ {a})" - by(simp add:insert_absorb[OF \a:A\]) - also have "\ = setsum f (A-{a}) + setsum f {a}" + from assms(3) obtain a where a: "a \ A" "f a < g a" by blast + have "setsum f A = setsum f ((A - {a}) \ {a})" + by(simp add: insert_absorb[OF \a \ A\]) + also have "\ = setsum f (A - {a}) + setsum f {a}" using \finite A\ by(subst setsum.union_disjoint) auto - also have "setsum f (A-{a}) \ setsum g (A-{a})" - by(rule setsum_mono)(simp add: assms(2)) - also have "setsum f {a} < setsum g {a}" using a by simp - also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \ {a})" - using \finite A\ by(subst setsum.union_disjoint[symmetric]) auto - also have "\ = setsum g A" by(simp add:insert_absorb[OF \a:A\]) - finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono) + also have "setsum f (A - {a}) \ setsum g (A - {a})" + by (rule setsum_mono) (simp add: assms(2)) + also from a have "setsum f {a} < setsum g {a}" by simp + also have "setsum g (A - {a}) + setsum g {a} = setsum g((A - {a}) \ {a})" + using \finite A\ by (subst setsum.union_disjoint[symmetric]) auto + also have "\ = setsum g A" by (simp add: insert_absorb[OF \a \ A\]) + finally show ?thesis + by (auto simp add: add_right_mono add_strict_left_mono) qed lemma setsum_mono_inv: @@ -609,51 +629,67 @@ assumes i: "i \ I" assumes I: "finite I" shows "f i = g i" -proof(rule ccontr) - assume "f i \ g i" +proof (rule ccontr) + assume "\ ?thesis" with le[OF i] have "f i < g i" by simp - hence "\i\I. f i < g i" using i .. - from setsum_strict_mono_ex1[OF I _ this] le have "setsum f I < setsum g I" by blast + with i have "\i\I. f i < g i" .. + from setsum_strict_mono_ex1[OF I _ this] le have "setsum f I < setsum g I" + by blast with eq show False by simp qed -lemma setsum_negf: "(\x\A. - f x::'a::ab_group_add) = - (\x\A. f x)" +lemma setsum_negf: "(\x\A. - f x) = - (\x\A. f x)" + for f :: "'b \ 'a::ab_group_add" proof (cases "finite A") - case True thus ?thesis by (induct set: finite) auto + case True + then show ?thesis by (induct set: finite) auto next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed -lemma setsum_subtractf: "(\x\A. f x - g x::'a::ab_group_add) = (\x\A. f x) - (\x\A. g x)" +lemma setsum_subtractf: "(\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" + for f g :: "'b \'a::ab_group_add" using setsum.distrib [of f "- g" A] by (simp add: setsum_negf) lemma setsum_subtractf_nat: - "(\x. x \ A \ g x \ f x) \ (\x\A. f x - g x::nat) = (\x\A. f x) - (\x\A. g x)" - by (induction A rule: infinite_finite_induct) (auto simp: setsum_mono) + "(\x. x \ A \ g x \ f x) \ (\x\A. f x - g x) = (\x\A. f x) - (\x\A. g x)" + for f g :: "'a \ nat" + by (induct A rule: infinite_finite_induct) (auto simp: setsum_mono) -lemma (in ordered_comm_monoid_add) setsum_nonneg: +context ordered_comm_monoid_add +begin + +lemma setsum_nonneg: assumes nn: "\x\A. 0 \ f x" shows "0 \ setsum f A" proof (cases "finite A") - case True thus ?thesis using nn + case True + then show ?thesis + using nn proof induct - case empty then show ?case by simp + case empty + then show ?case by simp next case (insert x F) then have "0 + 0 \ f x + setsum f F" by (blast intro: add_mono) with insert show ?case by simp qed next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed -lemma (in ordered_comm_monoid_add) setsum_nonpos: +lemma setsum_nonpos: assumes np: "\x\A. f x \ 0" shows "setsum f A \ 0" proof (cases "finite A") - case True thus ?thesis using np + case True + then show ?thesis + using np proof induct - case empty then show ?case by simp + case empty + then show ?case by simp next case (insert x F) then have "f x + setsum f F \ 0 + 0" by (blast intro: add_mono) @@ -663,232 +699,259 @@ case False thus ?thesis by simp qed -lemma (in ordered_comm_monoid_add) setsum_nonneg_eq_0_iff: +lemma setsum_nonneg_eq_0_iff: "finite A \ \x\A. 0 \ f x \ setsum f A = 0 \ (\x\A. f x = 0)" - by (induct set: finite, simp) (simp add: add_nonneg_eq_0_iff setsum_nonneg) + by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff setsum_nonneg) -lemma (in ordered_comm_monoid_add) setsum_nonneg_0: +lemma setsum_nonneg_0: "finite s \ (\i. i \ s \ f i \ 0) \ (\ i \ s. f i) = 0 \ i \ s \ f i = 0" by (simp add: setsum_nonneg_eq_0_iff) -lemma (in ordered_comm_monoid_add) setsum_nonneg_leq_bound: +lemma setsum_nonneg_leq_bound: assumes "finite s" "\i. i \ s \ f i \ 0" "(\i \ s. f i) = B" "i \ s" shows "f i \ B" proof - - have "f i \ f i + (\i \ s - {i}. f i)" - using assms by (intro add_increasing2 setsum_nonneg) auto + from assms have "f i \ f i + (\i \ s - {i}. f i)" + by (intro add_increasing2 setsum_nonneg) auto also have "\ = B" using setsum.remove[of s i f] assms by simp finally show ?thesis by auto qed -lemma (in ordered_comm_monoid_add) setsum_mono2: - assumes fin: "finite B" and sub: "A \ B" and nn: "\b. b \ B-A \ 0 \ f b" +lemma setsum_mono2: + assumes fin: "finite B" + and sub: "A \ B" + and nn: "\b. b \ B-A \ 0 \ f b" shows "setsum f A \ setsum f B" proof - have "setsum f A \ setsum f A + setsum f (B-A)" by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def) - also have "\ = setsum f (A \ (B-A))" using fin finite_subset[OF sub fin] - by (simp add: setsum.union_disjoint del:Un_Diff_cancel) - also have "A \ (B-A) = B" using sub by blast + also from fin finite_subset[OF sub fin] have "\ = setsum f (A \ (B-A))" + by (simp add: setsum.union_disjoint del: Un_Diff_cancel) + also from sub have "A \ (B-A) = B" by blast finally show ?thesis . qed -lemma (in ordered_comm_monoid_add) setsum_le_included: +lemma setsum_le_included: assumes "finite s" "finite t" and "\y\t. 0 \ g y" "(\x\s. \y\t. i y = x \ f x \ g y)" shows "setsum f s \ setsum g t" proof - have "setsum f s \ setsum (\y. setsum g {x. x\t \ i x = y}) s" proof (rule setsum_mono) - fix y assume "y \ s" + fix y + assume "y \ s" with assms obtain z where z: "z \ t" "y = i z" "f y \ g z" by auto with assms show "f y \ setsum g {x \ t. i x = y}" (is "?A y \ ?B y") using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro] by (auto intro!: setsum_mono2) qed - also have "... \ setsum (\y. setsum g {x. x\t \ i x = y}) (i ` t)" + also have "\ \ setsum (\y. setsum g {x. x\t \ i x = y}) (i ` t)" using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg) - also have "... \ setsum g t" + also have "\ \ setsum g t" using assms by (auto simp: setsum_image_gen[symmetric]) finally show ?thesis . qed -lemma (in ordered_comm_monoid_add) setsum_mono3: - "finite B \ A \ B \ \x\B - A. 0 \ f x \ setsum f A \ setsum f B" +lemma setsum_mono3: "finite B \ A \ B \ \x\B - A. 0 \ f x \ setsum f A \ setsum f B" by (rule setsum_mono2) auto +end + lemma (in canonically_ordered_monoid_add) setsum_eq_0_iff [simp]: "finite F \ (setsum f F = 0) = (\a\F. f a = 0)" by (intro ballI setsum_nonneg_eq_0_iff zero_le) lemma setsum_right_distrib: - fixes f :: "'a => ('b::semiring_0)" - shows "r * setsum f A = setsum (%n. r * f n) A" + fixes f :: "'a \ 'b::semiring_0" + shows "r * setsum f A = setsum (\n. r * f n) A" proof (cases "finite A") case True - thus ?thesis + then show ?thesis proof induct - case empty thus ?case by simp + case empty + then show ?case by simp next - case (insert x A) thus ?case by (simp add: distrib_left) + case insert + then show ?case by (simp add: distrib_left) qed next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed -lemma setsum_left_distrib: - "setsum f A * (r::'a::semiring_0) = (\n\A. f n * r)" +lemma setsum_left_distrib: "setsum f A * r = (\n\A. f n * r)" + for r :: "'a::semiring_0" proof (cases "finite A") case True then show ?thesis proof induct - case empty thus ?case by simp + case empty + then show ?case by simp next - case (insert x A) thus ?case by (simp add: distrib_right) + case insert + then show ?case by (simp add: distrib_right) + qed +next + case False + then show ?thesis by simp +qed + +lemma setsum_divide_distrib: "setsum f A / r = (\n\A. f n / r)" + for r :: "'a::field" +proof (cases "finite A") + case True + then show ?thesis + proof induct + case empty + then show ?case by simp + next + case insert + then show ?case by (simp add: add_divide_distrib) qed next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed -lemma setsum_divide_distrib: - "setsum f A / (r::'a::field) = (\n\A. f n / r)" +lemma setsum_abs[iff]: "\setsum f A\ \ setsum (\i. \f i\) A" + for f :: "'a \ 'b::ordered_ab_group_add_abs" +proof (cases "finite A") + case True + then show ?thesis + proof induct + case empty + then show ?case by simp + next + case insert + then show ?case by (auto intro: abs_triangle_ineq order_trans) + qed +next + case False + then show ?thesis by simp +qed + +lemma setsum_abs_ge_zero[iff]: "0 \ setsum (\i. \f i\) A" + for f :: "'a \ 'b::ordered_ab_group_add_abs" + by (simp add: setsum_nonneg) + +lemma abs_setsum_abs[simp]: "\\a\A. \f a\\ = (\a\A. \f a\)" + for f :: "'a \ 'b::ordered_ab_group_add_abs" proof (cases "finite A") case True then show ?thesis proof induct - case empty thus ?case by simp - next - case (insert x A) thus ?case by (simp add: add_divide_distrib) - qed -next - case False thus ?thesis by simp -qed - -lemma setsum_abs[iff]: - fixes f :: "'a => ('b::ordered_ab_group_add_abs)" - shows "\setsum f A\ \ setsum (%i. \f i\) A" -proof (cases "finite A") - case True - thus ?thesis - proof induct - case empty thus ?case by simp - next - case (insert x A) - thus ?case by (auto intro: abs_triangle_ineq order_trans) - qed -next - case False thus ?thesis by simp -qed - -lemma setsum_abs_ge_zero[iff]: - fixes f :: "'a => ('b::ordered_ab_group_add_abs)" - shows "0 \ setsum (%i. \f i\) A" - by (simp add: setsum_nonneg) - -lemma abs_setsum_abs[simp]: - fixes f :: "'a => ('b::ordered_ab_group_add_abs)" - shows "\\a\A. \f a\\ = (\a\A. \f a\)" -proof (cases "finite A") - case True - thus ?thesis - proof induct - case empty thus ?case by simp + case empty + then show ?case by simp next case (insert a A) - hence "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp - also have "\ = \\f a\ + \\a\A. \f a\\\" using insert by simp - also have "\ = \f a\ + \\a\A. \f a\\" - by (simp del: abs_of_nonneg) - also have "\ = (\a\insert a A. \f a\)" using insert by simp + then have "\\a\insert a A. \f a\\ = \\f a\ + (\a\A. \f a\)\" by simp + also from insert have "\ = \\f a\ + \\a\A. \f a\\\" by simp + also have "\ = \f a\ + \\a\A. \f a\\" by (simp del: abs_of_nonneg) + also from insert have "\ = (\a\insert a A. \f a\)" by simp finally show ?case . qed next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed -lemma setsum_diff1_ring: assumes "finite A" "a \ A" - shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)" +lemma setsum_diff1_ring: + fixes f :: "'b \ 'a::ring" + assumes "finite A" "a \ A" + shows "setsum f (A - {a}) = setsum f A - (f a)" unfolding setsum.remove [OF assms] by auto lemma setsum_product: - fixes f :: "'a => ('b::semiring_0)" + fixes f :: "'a \ 'b::semiring_0" shows "setsum f A * setsum g B = (\i\A. \j\B. f i * g j)" by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum.commute) lemma setsum_mult_setsum_if_inj: -fixes f :: "'a => ('b::semiring_0)" -shows "inj_on (%(a,b). f a * g b) (A \ B) ==> - setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}" -by(auto simp: setsum_product setsum.cartesian_product - intro!: setsum.reindex_cong[symmetric]) + fixes f :: "'a \ 'b::semiring_0" + shows "inj_on (\(a, b). f a * g b) (A \ B) \ + setsum f A * setsum g B = setsum id {f a * g b |a b. a \ A \ b \ B}" + by(auto simp: setsum_product setsum.cartesian_product intro!: setsum.reindex_cong[symmetric]) -lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" -apply (case_tac "finite A") - prefer 2 apply simp -apply (erule rev_mp) -apply (erule finite_induct, auto) -done +lemma setsum_SucD: + assumes "setsum f A = Suc n" + shows "\a\A. 0 < f a" +proof (cases "finite A") + case True + from this assms show ?thesis by induct auto +next + case False + with assms show ?thesis by simp +qed -lemma setsum_eq_Suc0_iff: "finite A \ - setsum f A = Suc 0 \ (EX a:A. f a = Suc 0 & (ALL b:A. a\b \ f b = 0))" -apply(erule finite_induct) -apply (auto simp add:add_is_1) -done +lemma setsum_eq_Suc0_iff: + assumes "finite A" + shows "setsum f A = Suc 0 \ (\a\A. f a = Suc 0 \ (\b\A. a \ b \ f b = 0))" + using assms by induct (auto simp add:add_is_1) lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]] -lemma setsum_Un_nat: "finite A ==> finite B ==> - (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" +lemma setsum_Un_nat: + "finite A \ finite B \ setsum f (A \ B) = setsum f A + setsum f B - setsum f (A \ B)" + for f :: "'a \ nat" \ \For the natural numbers, we have subtraction.\ -by (subst setsum.union_inter [symmetric], auto simp add: algebra_simps) + by (subst setsum.union_inter [symmetric]) (auto simp: algebra_simps) -lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) = - (if a:A then setsum f A - f a else setsum f A)" -apply (case_tac "finite A") - prefer 2 apply simp -apply (erule finite_induct) - apply (auto simp add: insert_Diff_if) -apply (drule_tac a = a in mk_disjoint_insert, auto) -done +lemma setsum_diff1_nat: "setsum f (A - {a}) = (if a \ A then setsum f A - f a else setsum f A)" + for f :: "'a \ nat" +proof (cases "finite A") + case True + then show ?thesis + apply induct + apply (auto simp: insert_Diff_if) + apply (drule mk_disjoint_insert) + apply auto + done +next + case False + then show ?thesis by simp +qed lemma setsum_diff_nat: -assumes "finite B" and "B \ A" -shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" -using assms + fixes f :: "'a \ nat" + assumes "finite B" and "B \ A" + shows "setsum f (A - B) = setsum f A - setsum f B" + using assms proof induct - show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp + case empty + then show ?case by simp next - fix F x assume finF: "finite F" and xnotinF: "x \ F" - and xFinA: "insert x F \ A" - and IH: "F \ A \ setsum f (A - F) = setsum f A - setsum f F" - from xnotinF xFinA have xinAF: "x \ (A - F)" by simp - from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x" + case (insert x F) + note IH = \F \ A \ setsum f (A - F) = setsum f A - setsum f F\ + from \x \ F\ \insert x F \ A\ have "x \ A - F" by simp + then have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x" by (simp add: setsum_diff1_nat) - from xFinA have "F \ A" by simp + from \insert x F \ A\ have "F \ A" by simp with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x" by simp - from xnotinF have "A - insert x F = (A - F) - {x}" by auto + from \x \ F\ have "A - insert x F = (A - F) - {x}" by auto with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x" by simp - from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp + from \finite F\ \x \ F\ have "setsum f (insert x F) = setsum f F + f x" + by simp with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp - thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp + then show ?case by simp qed lemma setsum_comp_morphism: assumes "h 0 = 0" and "\x y. h (x + y) = h x + h y" shows "setsum (h \ g) A = h (setsum g A)" proof (cases "finite A") - case False then show ?thesis by (simp add: assms) + case False + then show ?thesis by (simp add: assms) next - case True then show ?thesis by (induct A) (simp_all add: assms) + case True + then show ?thesis by (induct A) (simp_all add: assms) qed -lemma (in comm_semiring_1) dvd_setsum: - "(\a. a \ A \ d dvd f a) \ d dvd setsum f A" +lemma (in comm_semiring_1) dvd_setsum: "(\a. a \ A \ d dvd f a) \ d dvd setsum f A" by (induct A rule: infinite_finite_induct) simp_all lemma (in ordered_comm_monoid_add) setsum_pos: @@ -908,17 +971,18 @@ lemma setsum_cong_Suc: assumes "0 \ A" "\x. Suc x \ A \ f (Suc x) = g (Suc x)" - shows "setsum f A = setsum g A" + shows "setsum f A = setsum g A" proof (rule setsum.cong) - fix x assume "x \ A" - with assms(1) show "f x = g x" by (cases x) (auto intro!: assms(2)) + fix x + assume "x \ A" + with assms(1) show "f x = g x" + by (cases x) (auto intro!: assms(2)) qed simp_all subsubsection \Cardinality as special case of @{const setsum}\ -lemma card_eq_setsum: - "card A = setsum (\x. 1) A" +lemma card_eq_setsum: "card A = setsum (\x. 1) A" proof - have "plus \ (\_. Suc 0) = (\_. Suc)" by (simp add: fun_eq_iff) @@ -926,45 +990,53 @@ by (rule arg_cong) then have "Finite_Set.fold (plus \ (\_. Suc 0)) 0 A = Finite_Set.fold (\_. Suc) 0 A" by (blast intro: fun_cong) - then show ?thesis by (simp add: card.eq_fold setsum.eq_fold) + then show ?thesis + by (simp add: card.eq_fold setsum.eq_fold) qed -lemma setsum_constant [simp]: - "(\x \ A. y) = of_nat (card A) * y" -apply (cases "finite A") -apply (erule finite_induct) -apply (auto simp add: algebra_simps) -done +lemma setsum_constant [simp]: "(\x \ A. y) = of_nat (card A) * y" +proof (cases "finite A") + case True + then show ?thesis by induct (auto simp: algebra_simps) +next + case False + then show ?thesis by simp +qed lemma setsum_Suc: "setsum (\x. Suc(f x)) A = setsum f A + card A" - using setsum.distrib[of f "\_. 1" A] - by simp + using setsum.distrib[of f "\_. 1" A] by simp lemma setsum_bounded_above: - assumes le: "\i. i\A \ f i \ (K::'a::{semiring_1, ordered_comm_monoid_add})" + fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" + assumes le: "\i. i\A \ f i \ K" shows "setsum f A \ of_nat (card A) * K" proof (cases "finite A") case True - thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp + then show ?thesis + using le setsum_mono[where K=A and g = "\x. K"] by simp next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed lemma setsum_bounded_above_strict: - assumes "\i. i\A \ f i < (K::'a::{ordered_cancel_comm_monoid_add,semiring_1})" - "card A > 0" + fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}" + assumes "\i. i\A \ f i < K" "card A > 0" shows "setsum f A < of_nat (card A) * K" -using assms setsum_strict_mono[where A=A and g = "%x. K"] -by (simp add: card_gt_0_iff) + using assms setsum_strict_mono[where A=A and g = "\x. K"] + by (simp add: card_gt_0_iff) lemma setsum_bounded_below: - assumes le: "\i. i\A \ (K::'a::{semiring_1, ordered_comm_monoid_add}) \ f i" + fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}" + assumes le: "\i. i\A \ K \ f i" shows "of_nat (card A) * K \ setsum f A" proof (cases "finite A") case True - thus ?thesis using le setsum_mono[where K=A and f = "%x. K"] by simp + then show ?thesis + using le setsum_mono[where K=A and f = "%x. K"] by simp next - case False thus ?thesis by simp + case False + then show ?thesis by simp qed lemma card_UN_disjoint: @@ -972,24 +1044,26 @@ and "\i\I. \j\I. i \ j \ A i \ A j = {}" shows "card (UNION I A) = (\i\I. card(A i))" proof - - have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" by simp - with assms show ?thesis by (simp add: card_eq_setsum setsum.UNION_disjoint del: setsum_constant) + have "(\i\I. card (A i)) = (\i\I. \x\A i. 1)" + by simp + with assms show ?thesis + by (simp add: card_eq_setsum setsum.UNION_disjoint del: setsum_constant) qed lemma card_Union_disjoint: - "finite C ==> (ALL A:C. finite A) ==> - (ALL A:C. ALL B:C. A \ B --> A Int B = {}) - ==> card (\C) = setsum card C" -apply (frule card_UN_disjoint [of C id]) -apply simp_all -done + "finite C \ \A\C. finite A \ \A\C. \B\C. A \ B \ A \ B = {} \ + card (\C) = setsum card C" + by (frule card_UN_disjoint [of C id]) simp_all lemma setsum_multicount_gen: assumes "finite s" "finite t" "\j\t. (card {i\s. R i j} = k j)" - shows "setsum (\i. (card {j\t. R i j})) s = setsum k t" (is "?l = ?r") + shows "setsum (\i. (card {j\t. R i j})) s = setsum k t" + (is "?l = ?r") proof- - have "?l = setsum (\i. setsum (\x.1) {j\t. R i j}) s" by auto - also have "\ = ?r" unfolding setsum.commute_restrict [OF assms(1-2)] + have "?l = setsum (\i. setsum (\x.1) {j\t. R i j}) s" + by auto + also have "\ = ?r" + unfolding setsum.commute_restrict [OF assms(1-2)] using assms(3) by auto finally show ?thesis . qed @@ -998,17 +1072,18 @@ assumes "finite S" "finite T" "\j\T. (card {i\S. R i j} = k)" shows "setsum (\i. card {j\T. R i j}) S = k * card T" (is "?l = ?r") proof- - have "?l = setsum (\i. k) T" by (rule setsum_multicount_gen) (auto simp: assms) + have "?l = setsum (\i. k) T" + by (rule setsum_multicount_gen) (auto simp: assms) also have "\ = ?r" by (simp add: mult.commute) finally show ?thesis by auto qed + subsubsection \Cardinality of products\ lemma card_SigmaI [simp]: - "\ finite A; ALL a:A. finite (B a) \ - \ card (SIGMA x: A. B x) = (\a\A. card (B a))" -by(simp add: card_eq_setsum setsum.Sigma del:setsum_constant) + "finite A \ \a\A. finite (B a) \ card (SIGMA x: A. B x) = (\a\A. card (B a))" + by (simp add: card_eq_setsum setsum.Sigma del: setsum_constant) (* lemma SigmaI_insert: "y \ A ==> @@ -1016,12 +1091,12 @@ by auto *) -lemma card_cartesian_product: "card (A \ B) = card(A) * card(B)" +lemma card_cartesian_product: "card (A \ B) = card A * card B" by (cases "finite A \ finite B") (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) -lemma card_cartesian_product_singleton: "card({x} \ A) = card(A)" -by (simp add: card_cartesian_product) +lemma card_cartesian_product_singleton: "card ({x} \ A) = card A" + by (simp add: card_cartesian_product) subsection \Generalized product over a set\ @@ -1030,12 +1105,10 @@ begin sublocale setprod: comm_monoid_set times 1 -defines - setprod = setprod.F .. + defines setprod = setprod.F .. -abbreviation - Setprod ("\_" [1000] 999) where - "\A \ setprod (\x. x) A" +abbreviation Setprod ("\_" [1000] 999) + where "\A \ setprod (\x. x) A" end @@ -1058,22 +1131,26 @@ context comm_monoid_mult begin -lemma setprod_dvd_setprod: - "(\a. a \ A \ f a dvd g a) \ setprod f A dvd setprod g A" +lemma setprod_dvd_setprod: "(\a. a \ A \ f a dvd g a) \ setprod f A dvd setprod g A" proof (induct A rule: infinite_finite_induct) - case infinite then show ?case by (auto intro: dvdI) + case infinite + then show ?case by (auto intro: dvdI) +next + case empty + then show ?case by (auto intro: dvdI) next - case empty then show ?case by (auto intro: dvdI) -next - case (insert a A) then - have "f a dvd g a" and "setprod f A dvd setprod g A" by simp_all - then obtain r s where "g a = f a * r" and "setprod g A = setprod f A * s" by (auto elim!: dvdE) - then have "g a * setprod g A = f a * setprod f A * (r * s)" by (simp add: ac_simps) - with insert.hyps show ?case by (auto intro: dvdI) + case (insert a A) + then have "f a dvd g a" and "setprod f A dvd setprod g A" + by simp_all + then obtain r s where "g a = f a * r" and "setprod g A = setprod f A * s" + by (auto elim!: dvdE) + then have "g a * setprod g A = f a * setprod f A * (r * s)" + by (simp add: ac_simps) + with insert.hyps show ?case + by (auto intro: dvdI) qed -lemma setprod_dvd_setprod_subset: - "finite B \ A \ B \ setprod f A dvd setprod f B" +lemma setprod_dvd_setprod_subset: "finite B \ A \ B \ setprod f A dvd setprod f B" by (auto simp add: setprod.subset_diff ac_simps intro: dvdI) end @@ -1090,21 +1167,23 @@ proof - from \finite A\ have "setprod f (insert a (A - {a})) = f a * setprod f (A - {a})" by (intro setprod.insert) auto - also from \a \ A\ have "insert a (A - {a}) = A" by blast + also from \a \ A\ have "insert a (A - {a}) = A" + by blast finally have "setprod f A = f a * setprod f (A - {a})" . - with \b = f a\ show ?thesis by simp + with \b = f a\ show ?thesis + by simp qed -lemma dvd_setprodI [intro]: - assumes "finite A" and "a \ A" - shows "f a dvd setprod f A" - using assms by auto +lemma dvd_setprodI [intro]: "finite A \ a \ A \ f a dvd setprod f A" + by auto lemma setprod_zero: assumes "finite A" and "\a\A. f a = 0" shows "setprod f A = 0" -using assms proof (induct A) - case empty then show ?case by simp + using assms +proof (induct A) + case empty + then show ?case by simp next case (insert a A) then have "f a = 0 \ (\a\A. f a = 0)" by simp @@ -1126,71 +1205,73 @@ end lemma setprod_zero_iff [simp]: + fixes f :: "'b \ 'a::semidom" assumes "finite A" - shows "setprod f A = (0::'a::semidom) \ (\a\A. f a = 0)" + shows "setprod f A = 0 \ (\a\A. f a = 0)" using assms by (induct A) (auto simp: no_zero_divisors) lemma (in semidom_divide) setprod_diff1: assumes "finite A" and "f a \ 0" shows "setprod f (A - {a}) = (if a \ A then setprod f A div f a else setprod f A)" proof (cases "a \ A") - case True then show ?thesis by simp + case True + then show ?thesis by simp next - case False with assms show ?thesis - proof (induct A rule: finite_induct) - case empty then show ?case by simp + case False + with assms show ?thesis + proof induct + case empty + then show ?case by simp next case (insert b B) then show ?case proof (cases "a = b") - case True with insert show ?thesis by simp + case True + with insert show ?thesis by simp next - case False with insert have "a \ B" by simp + case False + with insert have "a \ B" by simp define C where "C = B - {a}" - with \finite B\ \a \ B\ - have *: "B = insert a C" "finite C" "a \ C" by auto - with insert show ?thesis by (auto simp add: insert_commute ac_simps) + with \finite B\ \a \ B\ have "B = insert a C" "finite C" "a \ C" + by auto + with insert show ?thesis + by (auto simp add: insert_commute ac_simps) qed qed qed -lemma setsum_zero_power [simp]: - fixes c :: "nat \ 'a::division_ring" - shows "(\i\A. c i * 0^i) = (if finite A \ 0 \ A then c 0 else 0)" -apply (cases "finite A") - by (induction A rule: finite_induct) auto +lemma setsum_zero_power [simp]: "(\i\A. c i * 0^i) = (if finite A \ 0 \ A then c 0 else 0)" + for c :: "nat \ 'a::division_ring" + by (induct A rule: infinite_finite_induct) auto lemma setsum_zero_power' [simp]: - fixes c :: "nat \ 'a::field" - shows "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" - using setsum_zero_power [of "\i. c i / d i" A] - by auto + "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" + for c :: "nat \ 'a::field" + using setsum_zero_power [of "\i. c i / d i" A] by auto lemma (in field) setprod_inversef: "finite A \ setprod (inverse \ f) A = inverse (setprod f A)" by (induct A rule: finite_induct) simp_all -lemma (in field) setprod_dividef: - "finite A \ (\x\A. f x / g x) = setprod f A / setprod g A" +lemma (in field) setprod_dividef: "finite A \ (\x\A. f x / g x) = setprod f A / setprod g A" using setprod_inversef [of A g] by (simp add: divide_inverse setprod.distrib) lemma setprod_Un: fixes f :: "'b \ 'a :: field" assumes "finite A" and "finite B" - and "\x\A \ B. f x \ 0" + and "\x\A \ B. f x \ 0" shows "setprod f (A \ B) = setprod f A * setprod f B / setprod f (A \ B)" proof - from assms have "setprod f A * setprod f B = setprod f (A \ B) * setprod f (A \ B)" by (simp add: setprod.union_inter [symmetric, of A B]) - with assms show ?thesis by simp + with assms show ?thesis + by simp qed -lemma (in linordered_semidom) setprod_nonneg: - "(\a\A. 0 \ f a) \ 0 \ setprod f A" +lemma (in linordered_semidom) setprod_nonneg: "(\a\A. 0 \ f a) \ 0 \ setprod f A" by (induct A rule: infinite_finite_induct) simp_all -lemma (in linordered_semidom) setprod_pos: - "(\a\A. 0 < f a) \ 0 < setprod f A" +lemma (in linordered_semidom) setprod_pos: "(\a\A. 0 < f a) \ 0 < setprod f A" by (induct A rule: infinite_finite_induct) simp_all lemma (in linordered_semidom) setprod_mono: @@ -1198,71 +1279,69 @@ by (induct A rule: infinite_finite_induct) (auto intro!: setprod_nonneg mult_mono) lemma (in linordered_semidom) setprod_mono_strict: - assumes"finite A" "\i\A. 0 \ f i \ f i < g i" "A \ {}" - shows "setprod f A < setprod g A" -using assms -apply (induct A rule: finite_induct) -apply (simp add: ) -apply (force intro: mult_strict_mono' setprod_nonneg) -done + assumes "finite A" "\i\A. 0 \ f i \ f i < g i" "A \ {}" + shows "setprod f A < setprod g A" + using assms +proof (induct A rule: finite_induct) + case empty + then show ?case by simp +next + case insert + then show ?case by (force intro: mult_strict_mono' setprod_nonneg) +qed -lemma (in linordered_field) abs_setprod: - "\setprod f A\ = (\x\A. \f x\)" +lemma (in linordered_field) abs_setprod: "\setprod f A\ = (\x\A. \f x\)" by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult) -lemma setprod_eq_1_iff [simp]: - "finite A \ setprod f A = 1 \ (\a\A. f a = (1::nat))" +lemma setprod_eq_1_iff [simp]: "finite A \ setprod f A = 1 \ (\a\A. f a = 1)" + for f :: "'a \ nat" by (induct A rule: finite_induct) simp_all -lemma setprod_pos_nat_iff [simp]: - "finite A \ setprod f A > 0 \ (\a\A. f a > (0::nat))" +lemma setprod_pos_nat_iff [simp]: "finite A \ setprod f A > 0 \ (\a\A. f a > 0)" + for f :: "'a \ nat" using setprod_zero_iff by (simp del: neq0_conv add: zero_less_iff_neq_zero) -lemma setprod_constant: - "(\x\ A. (y::'a::comm_monoid_mult)) = y ^ card A" +lemma setprod_constant: "(\x\ A. y) = y ^ card A" + for y :: "'a::comm_monoid_mult" by (induct A rule: infinite_finite_induct) simp_all -lemma setprod_power_distrib: - fixes f :: "'a \ 'b::comm_semiring_1" - shows "setprod f A ^ n = setprod (\x. (f x) ^ n) A" -proof (cases "finite A") - case True then show ?thesis - by (induct A rule: finite_induct) (auto simp add: power_mult_distrib) -next - case False then show ?thesis - by simp -qed +lemma setprod_power_distrib: "setprod f A ^ n = setprod (\x. (f x) ^ n) A" + for f :: "'a \ 'b::comm_semiring_1" + by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib) -lemma power_setsum: - "c ^ (\a\A. f a) = (\a\A. c ^ f a)" +lemma power_setsum: "c ^ (\a\A. f a) = (\a\A. c ^ f a)" by (induct A rule: infinite_finite_induct) (simp_all add: power_add) lemma setprod_gen_delta: - assumes fS: "finite S" - shows "setprod (\k. if k=a then b k else c) S = (if a \ S then (b a ::'a::comm_monoid_mult) * c^ (card S - 1) else c^ card S)" -proof- + fixes b :: "'b \ 'a::comm_monoid_mult" + assumes fin: "finite S" + shows "setprod (\k. if k = a then b k else c) S = + (if a \ S then b a * c ^ (card S - 1) else c ^ card S)" +proof - let ?f = "(\k. if k=a then b k else c)" - {assume a: "a \ S" - hence "\ k\ S. ?f k = c" by simp - hence ?thesis using a setprod_constant by simp } - moreover - {assume a: "a \ S" + show ?thesis + proof (cases "a \ S") + case False + then have "\ k\ S. ?f k = c" by simp + with False show ?thesis by (simp add: setprod_constant) + next + case True let ?A = "S - {a}" let ?B = "{a}" - have eq: "S = ?A \ ?B" using a by blast - have dj: "?A \ ?B = {}" by simp - from fS have fAB: "finite ?A" "finite ?B" by auto - have fA0:"setprod ?f ?A = setprod (\i. c) ?A" + from True have eq: "S = ?A \ ?B" by blast + have disjoint: "?A \ ?B = {}" by simp + from fin have fin': "finite ?A" "finite ?B" by auto + have f_A0: "setprod ?f ?A = setprod (\i. c) ?A" by (rule setprod.cong) auto - have cA: "card ?A = card S - 1" using fS a by auto - have fA1: "setprod ?f ?A = c ^ card ?A" - unfolding fA0 by (rule setprod_constant) + from fin True have card_A: "card ?A = card S - 1" by auto + have f_A1: "setprod ?f ?A = c ^ card ?A" + unfolding f_A0 by (rule setprod_constant) have "setprod ?f ?A * setprod ?f ?B = setprod ?f S" - using setprod.union_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]] + using setprod.union_disjoint[OF fin' disjoint, of ?f, unfolded eq[symmetric]] by simp - then have ?thesis using a cA - by (simp add: fA1 field_simps cong add: setprod.cong cong del: if_weak_cong)} - ultimately show ?thesis by blast + with True card_A show ?thesis + by (simp add: f_A1 field_simps cong add: setprod.cong cong del: if_weak_cong) + qed qed end diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Num.thy --- a/src/HOL/Num.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Num.thy Wed Aug 10 22:05:36 2016 +0200 @@ -6,7 +6,7 @@ section \Binary Numerals\ theory Num -imports BNF_Least_Fixpoint + imports BNF_Least_Fixpoint begin subsection \The \num\ type\ @@ -15,21 +15,24 @@ text \Increment function for type @{typ num}\ -primrec inc :: "num \ num" where - "inc One = Bit0 One" | - "inc (Bit0 x) = Bit1 x" | - "inc (Bit1 x) = Bit0 (inc x)" +primrec inc :: "num \ num" + where + "inc One = Bit0 One" + | "inc (Bit0 x) = Bit1 x" + | "inc (Bit1 x) = Bit0 (inc x)" text \Converting between type @{typ num} and type @{typ nat}\ -primrec nat_of_num :: "num \ nat" where - "nat_of_num One = Suc 0" | - "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | - "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" +primrec nat_of_num :: "num \ nat" + where + "nat_of_num One = Suc 0" + | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" + | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" -primrec num_of_nat :: "nat \ num" where - "num_of_nat 0 = One" | - "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" +primrec num_of_nat :: "nat \ num" + where + "num_of_nat 0 = One" + | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" lemma nat_of_num_pos: "0 < nat_of_num x" by (induct x) simp_all @@ -40,14 +43,10 @@ lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" by (induct x) simp_all -lemma num_of_nat_double: - "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" +lemma num_of_nat_double: "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" by (induct n) simp_all -text \ - Type @{typ num} is isomorphic to the strictly positive - natural numbers. -\ +text \Type @{typ num} is isomorphic to the strictly positive natural numbers.\ lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) @@ -68,10 +67,11 @@ shows "P x" proof - obtain n where n: "Suc n = nat_of_num x" - by (cases "nat_of_num x", simp_all add: nat_of_num_neq_0) + by (cases "nat_of_num x") (simp_all add: nat_of_num_neq_0) have "P (num_of_nat (Suc n))" proof (induct n) - case 0 show ?case using One by simp + case 0 + from One show ?case by simp next case (Suc n) then have "P (inc (num_of_nat (Suc n)))" by (rule inc) @@ -82,9 +82,9 @@ qed text \ - From now on, there are two possible models for @{typ num}: - as positive naturals (rule \num_induct\) - and as digit representation (rules \num.induct\, \num.cases\). + From now on, there are two possible models for @{typ num}: as positive + naturals (rule \num_induct\) and as digit representation (rules + \num.induct\, \num.cases\). \ @@ -93,17 +93,13 @@ instantiation num :: "{plus,times,linorder}" begin -definition [code del]: - "m + n = num_of_nat (nat_of_num m + nat_of_num n)" +definition [code del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)" -definition [code del]: - "m * n = num_of_nat (nat_of_num m * nat_of_num n)" +definition [code del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)" -definition [code del]: - "m \ n \ nat_of_num m \ nat_of_num n" +definition [code del]: "m \ n \ nat_of_num m \ nat_of_num n" -definition [code del]: - "m < n \ nat_of_num m < nat_of_num n" +definition [code del]: "m < n \ nat_of_num m < nat_of_num n" instance by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) @@ -137,8 +133,7 @@ "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" - by (simp_all add: num_eq_iff nat_of_num_add - nat_of_num_mult distrib_right distrib_left) + by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult distrib_right distrib_left) lemma eq_num_simps: "One = One \ True" @@ -175,9 +170,9 @@ by (auto simp add: less_eq_num_def less_num_def) lemma le_num_One_iff: "x \ num.One \ x = num.One" -by (simp add: antisym_conv) + by (simp add: antisym_conv) -text \Rules using \One\ and \inc\ as constructors\ +text \Rules using \One\ and \inc\ as constructors.\ lemma add_One: "x + One = inc x" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) @@ -191,22 +186,22 @@ lemma mult_inc: "x * inc y = x * y + x" by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) -text \The @{const num_of_nat} conversion\ +text \The @{const num_of_nat} conversion.\ -lemma num_of_nat_One: - "n \ 1 \ num_of_nat n = One" +lemma num_of_nat_One: "n \ 1 \ num_of_nat n = One" by (cases n) simp_all lemma num_of_nat_plus_distrib: "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" by (induct n) (auto simp add: add_One add_One_commute add_inc) -text \A double-and-decrement function\ +text \A double-and-decrement function.\ -primrec BitM :: "num \ num" where - "BitM One = One" | - "BitM (Bit0 n) = Bit1 (BitM n)" | - "BitM (Bit1 n) = Bit1 (Bit0 n)" +primrec BitM :: "num \ num" + where + "BitM One = One" + | "BitM (Bit0 n) = Bit1 (BitM n)" + | "BitM (Bit1 n) = Bit1 (Bit0 n)" lemma BitM_plus_one: "BitM n + One = Bit0 n" by (induct n) simp_all @@ -214,20 +209,22 @@ lemma one_plus_BitM: "One + BitM n = Bit0 n" unfolding add_One_commute BitM_plus_one .. -text \Squaring and exponentiation\ +text \Squaring and exponentiation.\ -primrec sqr :: "num \ num" where - "sqr One = One" | - "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | - "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" +primrec sqr :: "num \ num" + where + "sqr One = One" + | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" + | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" -primrec pow :: "num \ num \ num" where - "pow x One = x" | - "pow x (Bit0 y) = sqr (pow x y)" | - "pow x (Bit1 y) = sqr (pow x y) * x" +primrec pow :: "num \ num \ num" + where + "pow x One = x" + | "pow x (Bit0 y) = sqr (pow x y)" + | "pow x (Bit1 y) = sqr (pow x y) * x" lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" - by (induct x, simp_all add: algebra_simps nat_of_num_add) + by (induct x) (simp_all add: algebra_simps nat_of_num_add) lemma sqr_conv_mult: "sqr x = x * x" by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) @@ -243,32 +240,44 @@ class numeral = one + semigroup_add begin -primrec numeral :: "num \ 'a" where - numeral_One: "numeral One = 1" | - numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | - numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" +primrec numeral :: "num \ 'a" + where + numeral_One: "numeral One = 1" + | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" + | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" lemma numeral_code [code]: "numeral One = 1" "numeral (Bit0 n) = (let m = numeral n in m + m)" "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" by (simp_all add: Let_def) - + lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" - apply (induct x) - apply simp - apply (simp add: add.assoc [symmetric], simp add: add.assoc) - apply (simp add: add.assoc [symmetric], simp add: add.assoc) - done +proof (induct x) + case One + then show ?case by simp +next + case Bit0 + then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) +next + case Bit1 + then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) +qed lemma numeral_inc: "numeral (inc x) = numeral x + 1" proof (induct x) + case One + then show ?case by simp +next + case Bit0 + then show ?case by simp +next case (Bit1 x) have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case by (simp add: add.assoc) -qed simp_all +qed declare numeral.simps [simp del] @@ -320,9 +329,8 @@ subsection \Class-specific numeral rules\ -text \ - @{const numeral} is a morphism. -\ +text \@{const numeral} is a morphism.\ + subsubsection \Structures with addition: class \numeral\\ @@ -331,7 +339,7 @@ lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) - (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) + (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) @@ -350,44 +358,43 @@ end -subsubsection \ - Structures with negation: class \neg_numeral\ -\ + +subsubsection \Structures with negation: class \neg_numeral\\ class neg_numeral = numeral + group_add begin -lemma uminus_numeral_One: - "- Numeral1 = - 1" +lemma uminus_numeral_One: "- Numeral1 = - 1" by (simp add: numeral_One) text \Numerals form an abelian subgroup.\ -inductive is_num :: "'a \ bool" where - "is_num 1" | - "is_num x \ is_num (- x)" | - "\is_num x; is_num y\ \ is_num (x + y)" +inductive is_num :: "'a \ bool" + where + "is_num 1" + | "is_num x \ is_num (- x)" + | "is_num x \ is_num y \ is_num (x + y)" lemma is_num_numeral: "is_num (numeral k)" - by (induct k, simp_all add: numeral.simps is_num.intros) + by (induct k) (simp_all add: numeral.simps is_num.intros) -lemma is_num_add_commute: - "\is_num x; is_num y\ \ x + y = y + x" +lemma is_num_add_commute: "is_num x \ is_num y \ x + y = y + x" apply (induct x rule: is_num.induct) - apply (induct y rule: is_num.induct) - apply simp - apply (rule_tac a=x in add_left_imp_eq) - apply (rule_tac a=x in add_right_imp_eq) + apply (induct y rule: is_num.induct) + apply simp + apply (rule_tac a=x in add_left_imp_eq) + apply (rule_tac a=x in add_right_imp_eq) + apply (simp add: add.assoc) + apply (simp add: add.assoc [symmetric]) + apply (simp add: add.assoc) + apply (rule_tac a=x in add_left_imp_eq) + apply (rule_tac a=x in add_right_imp_eq) + apply (simp add: add.assoc) apply (simp add: add.assoc) - apply (simp add: add.assoc [symmetric], simp add: add.assoc) - apply (rule_tac a=x in add_left_imp_eq) - apply (rule_tac a=x in add_right_imp_eq) - apply (simp add: add.assoc) - apply (simp add: add.assoc, simp add: add.assoc [symmetric]) + apply (simp add: add.assoc [symmetric]) done -lemma is_num_add_left_commute: - "\is_num x; is_num y\ \ x + (y + z) = y + (x + z)" +lemma is_num_add_left_commute: "is_num x \ is_num y \ x + (y + z) = y + (x + z)" by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = @@ -395,12 +402,17 @@ is_num.intros is_num_numeral minus_add -definition dbl :: "'a \ 'a" where "dbl x = x + x" -definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" -definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" +definition dbl :: "'a \ 'a" + where "dbl x = x + x" + +definition dbl_inc :: "'a \ 'a" + where "dbl_inc x = x + x + 1" -definition sub :: "num \ num \ 'a" where - "sub k l = numeral k - numeral l" +definition dbl_dec :: "'a \ 'a" + where "dbl_dec x = x + x - 1" + +definition sub :: "num \ num \ 'a" + where "sub k l = numeral k - numeral l" lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) @@ -419,7 +431,8 @@ "dbl_inc 1 = 3" "dbl_inc (- 1) = - 1" "dbl_inc (numeral k) = numeral (Bit1 k)" - by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) + by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps + del: add_uminus_conv_diff) lemma dbl_dec_simps [simp]: "dbl_dec (- numeral k) = - dbl_inc (numeral k)" @@ -447,7 +460,7 @@ "- numeral m + numeral n = sub n m" "- numeral m + - numeral n = - (numeral m + numeral n)" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize - del: add_uminus_conv_diff add: diff_conv_add_uminus) + del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_special: "1 + - numeral m = sub One m" @@ -460,7 +473,7 @@ "- 1 + 1 = 0" "- 1 + - 1 = - 2" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc - del: add_uminus_conv_diff add: diff_conv_add_uminus) + del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_simps: "numeral m - numeral n = sub m n" @@ -468,7 +481,7 @@ "- numeral m - numeral n = - numeral (m + n)" "- numeral m - - numeral n = sub n m" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize - del: add_uminus_conv_diff add: diff_conv_add_uminus) + del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_special: "1 - numeral n = sub One n" @@ -484,13 +497,12 @@ "1 - - 1 = 2" "- 1 - - 1 = 0" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc - del: add_uminus_conv_diff add: diff_conv_add_uminus) + del: add_uminus_conv_diff add: diff_conv_add_uminus) end -subsubsection \ - Structures with multiplication: class \semiring_numeral\ -\ + +subsubsection \Structures with multiplication: class \semiring_numeral\\ class semiring_numeral = semiring + monoid_mult begin @@ -498,25 +510,22 @@ subclass numeral .. lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" - apply (induct n rule: num_induct) - apply (simp add: numeral_One) - apply (simp add: mult_inc numeral_inc numeral_add distrib_left) - done + by (induct n rule: num_induct) + (simp_all add: numeral_One mult_inc numeral_inc numeral_add distrib_left) lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" by (rule numeral_mult [symmetric]) lemma mult_2: "2 * z = z + z" - unfolding one_add_one [symmetric] distrib_right by simp + by (simp add: one_add_one [symmetric] distrib_right) lemma mult_2_right: "z * 2 = z + z" - unfolding one_add_one [symmetric] distrib_left by simp + by (simp add: one_add_one [symmetric] distrib_left) end -subsubsection \ - Structures with a zero: class \semiring_1\ -\ + +subsubsection \Structures with a zero: class \semiring_1\\ context semiring_1 begin @@ -524,18 +533,17 @@ subclass semiring_numeral .. lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" - by (induct n, - simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) + by (induct n) (simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) end -lemma nat_of_num_numeral [code_abbrev]: - "nat_of_num = numeral" +lemma nat_of_num_numeral [code_abbrev]: "nat_of_num = numeral" proof fix n have "numeral n = nat_of_num n" by (induct n) (simp_all add: numeral.simps) - then show "nat_of_num n = numeral n" by simp + then show "nat_of_num n = numeral n" + by simp qed lemma nat_of_num_code [code]: @@ -544,16 +552,15 @@ "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" by (simp_all add: Let_def) -subsubsection \ - Equality: class \semiring_char_0\ -\ + +subsubsection \Equality: class \semiring_char_0\\ context semiring_char_0 begin lemma numeral_eq_iff: "numeral m = numeral n \ m = n" - unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] - of_nat_eq_iff num_eq_iff .. + by (simp only: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] + of_nat_eq_iff num_eq_iff) lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" by (rule numeral_eq_iff [of n One, unfolded numeral_One]) @@ -562,8 +569,7 @@ by (rule numeral_eq_iff [of One n, unfolded numeral_One]) lemma numeral_neq_zero: "numeral n \ 0" - unfolding of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] - by (simp add: nat_of_num_pos) + by (simp add: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] nat_of_num_pos) lemma zero_neq_numeral: "0 \ numeral n" unfolding eq_commute [of 0] by (rule numeral_neq_zero) @@ -577,9 +583,8 @@ end -subsubsection \ - Comparisons: class \linordered_semidom\ -\ + +subsubsection \Comparisons: class \linordered_semidom\\ text \Could be perhaps more general than here.\ @@ -589,15 +594,15 @@ lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" proof - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" - unfolding less_eq_num_def nat_of_num_numeral of_nat_le_iff .. + by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) then show ?thesis by simp qed lemma one_le_numeral: "1 \ numeral n" -using numeral_le_iff [of One n] by (simp add: numeral_One) + using numeral_le_iff [of One n] by (simp add: numeral_One) lemma numeral_le_one_iff: "numeral n \ 1 \ n \ One" -using numeral_le_iff [of n One] by (simp add: numeral_One) + using numeral_le_iff [of n One] by (simp add: numeral_One) lemma numeral_less_iff: "numeral m < numeral n \ m < n" proof - @@ -647,30 +652,33 @@ not_numeral_less_zero lemma min_0_1 [simp]: - fixes min' :: "'a \ 'a \ 'a" defines "min' \ min" shows - "min' 0 1 = 0" - "min' 1 0 = 0" - "min' 0 (numeral x) = 0" - "min' (numeral x) 0 = 0" - "min' 1 (numeral x) = 1" - "min' (numeral x) 1 = 1" -by(simp_all add: min'_def min_def le_num_One_iff) + fixes min' :: "'a \ 'a \ 'a" + defines "min' \ min" + shows + "min' 0 1 = 0" + "min' 1 0 = 0" + "min' 0 (numeral x) = 0" + "min' (numeral x) 0 = 0" + "min' 1 (numeral x) = 1" + "min' (numeral x) 1 = 1" + by (simp_all add: min'_def min_def le_num_One_iff) -lemma max_0_1 [simp]: - fixes max' :: "'a \ 'a \ 'a" defines "max' \ max" shows - "max' 0 1 = 1" - "max' 1 0 = 1" - "max' 0 (numeral x) = numeral x" - "max' (numeral x) 0 = numeral x" - "max' 1 (numeral x) = numeral x" - "max' (numeral x) 1 = numeral x" -by(simp_all add: max'_def max_def le_num_One_iff) +lemma max_0_1 [simp]: + fixes max' :: "'a \ 'a \ 'a" + defines "max' \ max" + shows + "max' 0 1 = 1" + "max' 1 0 = 1" + "max' 0 (numeral x) = numeral x" + "max' (numeral x) 0 = numeral x" + "max' 1 (numeral x) = numeral x" + "max' (numeral x) 1 = numeral x" + by (simp_all add: max'_def max_def le_num_One_iff) end -subsubsection \ - Multiplication and negation: class \ring_1\ -\ + +subsubsection \Multiplication and negation: class \ring_1\\ context ring_1 begin @@ -681,20 +689,18 @@ "- numeral m * - numeral n = numeral (m * n)" "- numeral m * numeral n = - numeral (m * n)" "numeral m * - numeral n = - numeral (m * n)" - unfolding mult_minus_left mult_minus_right - by (simp_all only: minus_minus numeral_mult) + by (simp_all only: mult_minus_left mult_minus_right minus_minus numeral_mult) lemma mult_minus1 [simp]: "- 1 * z = - z" - unfolding numeral.simps mult_minus_left by simp + by (simp add: numeral.simps) lemma mult_minus1_right [simp]: "z * - 1 = - z" - unfolding numeral.simps mult_minus_right by simp + by (simp add: numeral.simps) end -subsubsection \ - Equality using \iszero\ for rings with non-zero characteristic -\ + +subsubsection \Equality using \iszero\ for rings with non-zero characteristic\ context ring_1 begin @@ -717,23 +723,22 @@ lemma not_iszero_neg_Numeral1: "\ iszero (- Numeral1)" by (simp add: numeral_One) -lemma iszero_neg_numeral [simp]: - "iszero (- numeral w) \ iszero (numeral w)" - unfolding iszero_def - by (rule neg_equal_0_iff_equal) +lemma iszero_neg_numeral [simp]: "iszero (- numeral w) \ iszero (numeral w)" + unfolding iszero_def by (rule neg_equal_0_iff_equal) lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" unfolding iszero_def by (rule eq_iff_diff_eq_0) -text \The \eq_numeral_iff_iszero\ lemmas are not declared -\[simp]\ by default, because for rings of characteristic zero, -better simp rules are possible. For a type like integers mod \n\, type-instantiated versions of these rules should be added to the -simplifier, along with a type-specific rule for deciding propositions -of the form \iszero (numeral w)\. +text \ + The \eq_numeral_iff_iszero\ lemmas are not declared \[simp]\ by default, + because for rings of characteristic zero, better simp rules are possible. + For a type like integers mod \n\, type-instantiated versions of these rules + should be added to the simplifier, along with a type-specific rule for + deciding propositions of the form \iszero (numeral w)\. -bh: Maybe it would not be so bad to just declare these as simp -rules anyway? I should test whether these rules take precedence over -the \ring_char_0\ rules in the simplifier. + bh: Maybe it would not be so bad to just declare these as simp rules anyway? + I should test whether these rules take precedence over the \ring_char_0\ + rules in the simplifier. \ lemma eq_numeral_iff_iszero: @@ -754,9 +759,8 @@ end -subsubsection \ - Equality and negation: class \ring_char_0\ -\ + +subsubsection \Equality and negation: class \ring_char_0\\ context ring_char_0 begin @@ -768,17 +772,16 @@ by simp lemma numeral_neq_neg_numeral: "numeral m \ - numeral n" - unfolding eq_neg_iff_add_eq_0 - by (simp add: numeral_plus_numeral) + by (simp add: eq_neg_iff_add_eq_0 numeral_plus_numeral) lemma neg_numeral_neq_numeral: "- numeral m \ numeral n" by (rule numeral_neq_neg_numeral [symmetric]) lemma zero_neq_neg_numeral: "0 \ - numeral n" - unfolding neg_0_equal_iff_equal by simp + by simp lemma neg_numeral_neq_zero: "- numeral n \ 0" - unfolding neg_equal_0_iff_equal by simp + by simp lemma one_neq_neg_numeral: "1 \ - numeral n" using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) @@ -786,36 +789,28 @@ lemma neg_numeral_neq_one: "- numeral n \ 1" using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) -lemma neg_one_neq_numeral: - "- 1 \ numeral n" +lemma neg_one_neq_numeral: "- 1 \ numeral n" using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) -lemma numeral_neq_neg_one: - "numeral n \ - 1" +lemma numeral_neq_neg_one: "numeral n \ - 1" using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) -lemma neg_one_eq_numeral_iff: - "- 1 = - numeral n \ n = One" +lemma neg_one_eq_numeral_iff: "- 1 = - numeral n \ n = One" using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) -lemma numeral_eq_neg_one_iff: - "- numeral n = - 1 \ n = One" +lemma numeral_eq_neg_one_iff: "- numeral n = - 1 \ n = One" using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) -lemma neg_one_neq_zero: - "- 1 \ 0" +lemma neg_one_neq_zero: "- 1 \ 0" by simp -lemma zero_neq_neg_one: - "0 \ - 1" +lemma zero_neq_neg_one: "0 \ - 1" by simp -lemma neg_one_neq_one: - "- 1 \ 1" +lemma neg_one_neq_one: "- 1 \ 1" using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) -lemma one_neq_neg_one: - "1 \ - 1" +lemma one_neq_neg_one: "1 \ - 1" using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemmas eq_neg_numeral_simps [simp] = @@ -831,9 +826,7 @@ end -subsubsection \ - Structures with negation and order: class \linordered_idom\ -\ +subsubsection \Structures with negation and order: class \linordered_idom\\ context linordered_idom begin @@ -869,7 +862,7 @@ lemma not_numeral_le_neg_numeral: "\ numeral m \ - numeral n" by (simp only: not_le neg_numeral_less_numeral) - + lemma neg_numeral_less_one: "- numeral m < 1" by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) @@ -906,20 +899,16 @@ lemma not_neg_one_le_neg_numeral_iff: "\ - 1 \ - numeral m \ m \ One" by (cases m) simp_all -lemma sub_non_negative: - "sub n m \ 0 \ n \ m" +lemma sub_non_negative: "sub n m \ 0 \ n \ m" by (simp only: sub_def le_diff_eq) simp -lemma sub_positive: - "sub n m > 0 \ n > m" +lemma sub_positive: "sub n m > 0 \ n > m" by (simp only: sub_def less_diff_eq) simp -lemma sub_non_positive: - "sub n m \ 0 \ n \ m" +lemma sub_non_positive: "sub n m \ 0 \ n \ m" by (simp only: sub_def diff_le_eq) simp -lemma sub_negative: - "sub n m < 0 \ n < m" +lemma sub_negative: "sub n m < 0 \ n < m" by (simp only: sub_def diff_less_eq) simp lemmas le_neg_numeral_simps [simp] = @@ -963,9 +952,8 @@ end -subsubsection \ - Natural numbers -\ + +subsubsection \Natural numbers\ lemma Suc_1 [simp]: "Suc 1 = 2" unfolding Suc_eq_plus1 by (rule one_add_one) @@ -977,7 +965,7 @@ where [code del]: "pred_numeral k = numeral k - 1" lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" - unfolding pred_numeral_def by simp + by (simp add: pred_numeral_def) lemma eval_nat_numeral: "numeral One = Suc 0" @@ -989,8 +977,7 @@ "pred_numeral One = 0" "pred_numeral (Bit0 k) = numeral (BitM k)" "pred_numeral (Bit1 k) = numeral (Bit0 k)" - unfolding pred_numeral_def eval_nat_numeral - by (simp_all only: diff_Suc_Suc diff_0) + by (simp_all only: pred_numeral_def eval_nat_numeral diff_Suc_Suc diff_0) lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (simp add: eval_nat_numeral) @@ -1001,12 +988,11 @@ lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" by (simp only: numeral_One One_nat_def) -lemma Suc_nat_number_of_add: - "Suc (numeral v + n) = numeral (v + One) + n" +lemma Suc_nat_number_of_add: "Suc (numeral v + n) = numeral (v + One) + n" by simp -(*Maps #n to n for n = 1, 2*) -lemmas numerals = numeral_One [where 'a=nat] numeral_2_eq_2 +lemma numerals: "Numeral1 = (1::nat)" "2 = Suc (Suc 0)" + by (rule numeral_One) (rule numeral_2_eq_2) text \Comparisons involving @{term Suc}.\ @@ -1034,26 +1020,21 @@ lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" by (simp add: numeral_eq_Suc) -lemma max_Suc_numeral [simp]: - "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" +lemma max_Suc_numeral [simp]: "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" by (simp add: numeral_eq_Suc) -lemma max_numeral_Suc [simp]: - "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" +lemma max_numeral_Suc [simp]: "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" by (simp add: numeral_eq_Suc) -lemma min_Suc_numeral [simp]: - "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" +lemma min_Suc_numeral [simp]: "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" by (simp add: numeral_eq_Suc) -lemma min_numeral_Suc [simp]: - "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" +lemma min_numeral_Suc [simp]: "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" by (simp add: numeral_eq_Suc) text \For @{term case_nat} and @{term rec_nat}.\ -lemma case_nat_numeral [simp]: - "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" +lemma case_nat_numeral [simp]: "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" by (simp add: numeral_eq_Suc) lemma case_nat_add_eq_if [simp]: @@ -1061,21 +1042,18 @@ by (simp add: numeral_eq_Suc) lemma rec_nat_numeral [simp]: - "rec_nat a f (numeral v) = - (let pv = pred_numeral v in f pv (rec_nat a f pv))" + "rec_nat a f (numeral v) = (let pv = pred_numeral v in f pv (rec_nat a f pv))" by (simp add: numeral_eq_Suc Let_def) lemma rec_nat_add_eq_if [simp]: - "rec_nat a f (numeral v + n) = - (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" + "rec_nat a f (numeral v + n) = (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" by (simp add: numeral_eq_Suc Let_def) -text \Case analysis on @{term "n < 2"}\ - +text \Case analysis on @{term "n < 2"}.\ lemma less_2_cases: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) -text \Removal of Small Numerals: 0, 1 and (in additive positions) 2\ +text \Removal of Small Numerals: 0, 1 and (in additive positions) 2.\ text \bh: Are these rules really a good idea?\ lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" @@ -1085,7 +1063,6 @@ by simp text \Can be used to eliminate long strings of Sucs, but not by default.\ - lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp @@ -1099,12 +1076,10 @@ subclass field_char_0 .. -lemma half_gt_zero_iff: - "0 < a / 2 \ 0 < a" (is "?P \ ?Q") +lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" by (auto simp add: field_simps) -lemma half_gt_zero [simp]: - "0 < a \ 0 < a / 2" +lemma half_gt_zero [simp]: "0 < a \ 0 < a / 2" by (simp add: half_gt_zero_iff) end @@ -1124,50 +1099,52 @@ subsection \Setting up simprocs\ -lemma mult_numeral_1: "Numeral1 * a = (a::'a::semiring_numeral)" +lemma mult_numeral_1: "Numeral1 * a = a" + for a :: "'a::semiring_numeral" by simp -lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::semiring_numeral)" +lemma mult_numeral_1_right: "a * Numeral1 = a" + for a :: "'a::semiring_numeral" by simp -lemma divide_numeral_1: "a / Numeral1 = (a::'a::field)" +lemma divide_numeral_1: "a / Numeral1 = a" + for a :: "'a::field" by simp -lemma inverse_numeral_1: - "inverse Numeral1 = (Numeral1::'a::division_ring)" +lemma inverse_numeral_1: "inverse Numeral1 = (Numeral1::'a::division_ring)" by simp -text\Theorem lists for the cancellation simprocs. The use of a binary -numeral for 1 reduces the number of special cases.\ +text \ + Theorem lists for the cancellation simprocs. The use of a binary + numeral for 1 reduces the number of special cases. +\ lemma mult_1s: - fixes a :: "'a::semiring_numeral" - and b :: "'b::ring_1" - shows "Numeral1 * a = a" - "a * Numeral1 = a" - "- Numeral1 * b = - b" - "b * - Numeral1 = - b" + "Numeral1 * a = a" + "a * Numeral1 = a" + "- Numeral1 * b = - b" + "b * - Numeral1 = - b" + for a :: "'a::semiring_numeral" and b :: "'b::ring_1" by simp_all setup \ Reorient_Proc.add (fn Const (@{const_name numeral}, _) $ _ => true - | Const (@{const_name uminus}, _) $ (Const (@{const_name numeral}, _) $ _) => true - | _ => false) + | Const (@{const_name uminus}, _) $ (Const (@{const_name numeral}, _) $ _) => true + | _ => false) \ -simproc_setup reorient_numeral - ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc +simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") = + Reorient_Proc.proc -subsubsection \Simplification of arithmetic operations on integer constants.\ +subsubsection \Simplification of arithmetic operations on integer constants\ lemmas arith_special = (* already declared simp above *) add_numeral_special add_neg_numeral_special diff_numeral_special -(* rules already in simpset *) -lemmas arith_extra_simps = +lemmas arith_extra_simps = (* rules already in simpset *) numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right minus_zero diff_numeral_simps diff_0 diff_0_right @@ -1195,7 +1172,7 @@ lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult -text \Simplification of relational operations\ +text \Simplification of relational operations.\ lemmas eq_numeral_extra = zero_neq_one one_neq_zero @@ -1215,34 +1192,28 @@ unfolding Let_def .. declaration \ -let +let fun number_of ctxt T n = if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort numeral})) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; in K ( - Lin_Arith.add_simps (@{thms arith_simps} @ @{thms more_arith_simps} - @ @{thms rel_simps} - @ @{thms pred_numeral_simps} - @ @{thms arith_special numeral_One} - @ @{thms of_nat_simps}) - #> Lin_Arith.add_simps [@{thm Suc_numeral}, - @{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1}, - @{thm le_Suc_numeral}, @{thm le_numeral_Suc}, - @{thm less_Suc_numeral}, @{thm less_numeral_Suc}, - @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc}, - @{thm mult_Suc}, @{thm mult_Suc_right}, - @{thm of_nat_numeral}] + Lin_Arith.add_simps + @{thms arith_simps more_arith_simps rel_simps pred_numeral_simps + arith_special numeral_One of_nat_simps} + #> Lin_Arith.add_simps + @{thms Suc_numeral Let_numeral Let_neg_numeral Let_0 Let_1 + le_Suc_numeral le_numeral_Suc less_Suc_numeral less_numeral_Suc + Suc_eq_numeral eq_numeral_Suc mult_Suc mult_Suc_right of_nat_numeral} #> Lin_Arith.set_number_of number_of) end \ -subsubsection \Simplification of arithmetic when nested to the right.\ +subsubsection \Simplification of arithmetic when nested to the right\ -lemma add_numeral_left [simp]: - "numeral v + (numeral w + z) = (numeral(v + w) + z)" +lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: @@ -1261,7 +1232,7 @@ hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec -subsection \code module namespace\ +subsection \Code module namespace\ code_identifier code_module Num \ (SML) Arith and (OCaml) Arith and (Haskell) Arith diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Parity.thy --- a/src/HOL/Parity.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Parity.thy Wed Aug 10 22:05:36 2016 +0200 @@ -6,7 +6,7 @@ section \Parity in rings and semirings\ theory Parity -imports Nat_Transfer + imports Nat_Transfer begin subsection \Ring structures with parity and \even\/\odd\ predicates\ @@ -21,19 +21,15 @@ subclass semiring_numeral .. abbreviation even :: "'a \ bool" -where - "even a \ 2 dvd a" + where "even a \ 2 dvd a" abbreviation odd :: "'a \ bool" -where - "odd a \ \ 2 dvd a" + where "odd a \ \ 2 dvd a" -lemma even_zero [simp]: - "even 0" +lemma even_zero [simp]: "even 0" by (fact dvd_0_right) -lemma even_plus_one_iff [simp]: - "even (a + 1) \ odd a" +lemma even_plus_one_iff [simp]: "even (a + 1) \ odd a" by (auto simp add: dvd_add_right_iff intro: odd_even_add) lemma evenE [elim?]: @@ -53,13 +49,11 @@ with * have "a = 2 * c + 1" by simp with that show thesis . qed - -lemma even_times_iff [simp]: - "even (a * b) \ even a \ even b" + +lemma even_times_iff [simp]: "even (a * b) \ even a \ even b" by (auto dest: even_multD) -lemma even_numeral [simp]: - "even (numeral (Num.Bit0 n))" +lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" proof - have "even (2 * numeral n)" unfolding even_times_iff by simp @@ -69,8 +63,7 @@ unfolding numeral.simps . qed -lemma odd_numeral [simp]: - "odd (numeral (Num.Bit1 n))" +lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" proof assume "even (numeral (num.Bit1 n))" then have "even (numeral n + numeral n + 1)" @@ -79,22 +72,18 @@ unfolding mult_2 . then have "2 dvd numeral n * 2 + 1" by (simp add: ac_simps) - with dvd_add_times_triv_left_iff [of 2 "numeral n" 1] - have "2 dvd 1" - by simp + then have "2 dvd 1" + using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp then show False by simp qed -lemma even_add [simp]: - "even (a + b) \ (even a \ even b)" +lemma even_add [simp]: "even (a + b) \ (even a \ even b)" by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) -lemma odd_add [simp]: - "odd (a + b) \ (\ (odd a \ odd b))" +lemma odd_add [simp]: "odd (a + b) \ (\ (odd a \ odd b))" by simp -lemma even_power [simp]: - "even (a ^ n) \ even a \ n > 0" +lemma even_power [simp]: "even (a ^ n) \ even a \ n > 0" by (induct n) auto end @@ -104,12 +93,10 @@ subclass comm_ring_1 .. -lemma even_minus [simp]: - "even (- a) \ even a" +lemma even_minus [simp]: "even (- a) \ even a" by (fact dvd_minus_iff) -lemma even_diff [simp]: - "even (a - b) \ even (a + b)" +lemma even_diff [simp]: "even (a - b) \ even (a + b)" using even_add [of a "- b"] by simp end @@ -117,17 +104,14 @@ subsection \Instances for @{typ nat} and @{typ int}\ -lemma even_Suc_Suc_iff [simp]: - "2 dvd Suc (Suc n) \ 2 dvd n" +lemma even_Suc_Suc_iff [simp]: "2 dvd Suc (Suc n) \ 2 dvd n" using dvd_add_triv_right_iff [of 2 n] by simp -lemma even_Suc [simp]: - "2 dvd Suc n \ \ 2 dvd n" +lemma even_Suc [simp]: "2 dvd Suc n \ \ 2 dvd n" by (induct n) auto -lemma even_diff_nat [simp]: - fixes m n :: nat - shows "2 dvd (m - n) \ m < n \ 2 dvd (m + n)" +lemma even_diff_nat [simp]: "2 dvd (m - n) \ m < n \ 2 dvd (m + n)" + for m n :: nat proof (cases "n \ m") case True then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) @@ -137,8 +121,8 @@ next case False then show ?thesis by simp -qed - +qed + instance nat :: semiring_parity proof show "\ 2 dvd (1 :: nat)" @@ -165,7 +149,8 @@ then obtain r where "Suc n = 2 * r" .. moreover from * obtain s where "m * n = 2 * s" .. then have "2 * s + m = m * Suc n" by simp - ultimately have " 2 * s + m = 2 * (m * r)" by (simp add: algebra_simps) + ultimately have " 2 * s + m = 2 * (m * r)" + by (simp add: algebra_simps) then have "m = 2 * (m * r - s)" by simp then show "2 dvd m" .. qed @@ -176,13 +161,12 @@ by (cases n) simp_all qed -lemma odd_pos: - "odd (n :: nat) \ 0 < n" +lemma odd_pos: "odd n \ 0 < n" + for n :: nat by (auto elim: oddE) -lemma Suc_double_not_eq_double: - fixes m n :: nat - shows "Suc (2 * m) \ 2 * n" +lemma Suc_double_not_eq_double: "Suc (2 * m) \ 2 * n" + for m n :: nat proof assume "Suc (2 * m) = 2 * n" moreover have "odd (Suc (2 * m))" and "even (2 * n)" @@ -190,37 +174,34 @@ ultimately show False by simp qed -lemma double_not_eq_Suc_double: - fixes m n :: nat - shows "2 * m \ Suc (2 * n)" +lemma double_not_eq_Suc_double: "2 * m \ Suc (2 * n)" + for m n :: nat using Suc_double_not_eq_double [of n m] by simp -lemma even_diff_iff [simp]: - fixes k l :: int - shows "2 dvd (k - l) \ 2 dvd (k + l)" +lemma even_diff_iff [simp]: "2 dvd (k - l) \ 2 dvd (k + l)" + for k l :: int using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: mult_2_right) -lemma even_abs_add_iff [simp]: - fixes k l :: int - shows "2 dvd (\k\ + l) \ 2 dvd (k + l)" +lemma even_abs_add_iff [simp]: "2 dvd (\k\ + l) \ 2 dvd (k + l)" + for k l :: int by (cases "k \ 0") (simp_all add: ac_simps) -lemma even_add_abs_iff [simp]: - fixes k l :: int - shows "2 dvd (k + \l\) \ 2 dvd (k + l)" +lemma even_add_abs_iff [simp]: "2 dvd (k + \l\) \ 2 dvd (k + l)" + for k l :: int using even_abs_add_iff [of l k] by (simp add: ac_simps) -lemma odd_Suc_minus_one [simp]: - "odd n \ Suc (n - Suc 0) = n" +lemma odd_Suc_minus_one [simp]: "odd n \ Suc (n - Suc 0) = n" by (auto elim: oddE) instance int :: ring_parity proof - show "\ 2 dvd (1 :: int)" by (simp add: dvd_int_unfold_dvd_nat) + show "\ 2 dvd (1 :: int)" + by (simp add: dvd_int_unfold_dvd_nat) +next fix k l :: int assume "\ 2 dvd k" moreover assume "\ 2 dvd l" - ultimately have "2 dvd (nat \k\ + nat \l\)" + ultimately have "2 dvd (nat \k\ + nat \l\)" by (auto simp add: dvd_int_unfold_dvd_nat intro: odd_even_add) then have "2 dvd (\k\ + \l\)" by (simp add: dvd_int_unfold_dvd_nat nat_add_distrib) @@ -237,12 +218,10 @@ then show "\l. k = l + 1" .. qed -lemma even_int_iff [simp]: - "even (int n) \ even n" +lemma even_int_iff [simp]: "even (int n) \ even n" by (simp add: dvd_int_iff) -lemma even_nat_iff: - "0 \ k \ even (nat k) \ even k" +lemma even_nat_iff: "0 \ k \ even (nat k) \ even k" by (simp add: even_int_iff [symmetric]) @@ -251,58 +230,47 @@ context ring_1 begin -lemma power_minus_even [simp]: - "even n \ (- a) ^ n = a ^ n" +lemma power_minus_even [simp]: "even n \ (- a) ^ n = a ^ n" by (auto elim: evenE) -lemma power_minus_odd [simp]: - "odd n \ (- a) ^ n = - (a ^ n)" +lemma power_minus_odd [simp]: "odd n \ (- a) ^ n = - (a ^ n)" by (auto elim: oddE) -lemma neg_one_even_power [simp]: - "even n \ (- 1) ^ n = 1" +lemma neg_one_even_power [simp]: "even n \ (- 1) ^ n = 1" by simp -lemma neg_one_odd_power [simp]: - "odd n \ (- 1) ^ n = - 1" +lemma neg_one_odd_power [simp]: "odd n \ (- 1) ^ n = - 1" by simp -end +end context linordered_idom begin -lemma zero_le_even_power: - "even n \ 0 \ a ^ n" +lemma zero_le_even_power: "even n \ 0 \ a ^ n" by (auto elim: evenE) -lemma zero_le_odd_power: - "odd n \ 0 \ a ^ n \ 0 \ a" +lemma zero_le_odd_power: "odd n \ 0 \ a ^ n \ 0 \ a" by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) -lemma zero_le_power_eq: - "0 \ a ^ n \ even n \ odd n \ 0 \ a" +lemma zero_le_power_eq: "0 \ a ^ n \ even n \ odd n \ 0 \ a" by (auto simp add: zero_le_even_power zero_le_odd_power) - -lemma zero_less_power_eq: - "0 < a ^ n \ n = 0 \ even n \ a \ 0 \ odd n \ 0 < a" + +lemma zero_less_power_eq: "0 < a ^ n \ n = 0 \ even n \ a \ 0 \ odd n \ 0 < a" proof - have [simp]: "0 = a ^ n \ a = 0 \ n > 0" unfolding power_eq_0_iff [of a n, symmetric] by blast show ?thesis - unfolding less_le zero_le_power_eq by auto + unfolding less_le zero_le_power_eq by auto qed -lemma power_less_zero_eq [simp]: - "a ^ n < 0 \ odd n \ a < 0" +lemma power_less_zero_eq [simp]: "a ^ n < 0 \ odd n \ a < 0" unfolding not_le [symmetric] zero_le_power_eq by auto - -lemma power_le_zero_eq: - "a ^ n \ 0 \ n > 0 \ (odd n \ a \ 0 \ even n \ a = 0)" - unfolding not_less [symmetric] zero_less_power_eq by auto -lemma power_even_abs: - "even n \ \a\ ^ n = a ^ n" +lemma power_le_zero_eq: "a ^ n \ 0 \ n > 0 \ (odd n \ a \ 0 \ even n \ a = 0)" + unfolding not_less [symmetric] zero_less_power_eq by auto + +lemma power_even_abs: "even n \ \a\ ^ n = a ^ n" using power_abs [of a n] by (simp add: zero_le_even_power) lemma power_mono_even: @@ -310,30 +278,35 @@ shows "a ^ n \ b ^ n" proof - have "0 \ \a\" by auto - with \\a\ \ \b\\ - have "\a\ ^ n \ \b\ ^ n" by (rule power_mono) - with \even n\ show ?thesis by (simp add: power_even_abs) + with \\a\ \ \b\\ have "\a\ ^ n \ \b\ ^ n" + by (rule power_mono) + with \even n\ show ?thesis + by (simp add: power_even_abs) qed lemma power_mono_odd: assumes "odd n" and "a \ b" shows "a ^ n \ b ^ n" proof (cases "b < 0") - case True with \a \ b\ have "- b \ - a" and "0 \ - b" by auto - hence "(- b) ^ n \ (- a) ^ n" by (rule power_mono) + case True + with \a \ b\ have "- b \ - a" and "0 \ - b" by auto + then have "(- b) ^ n \ (- a) ^ n" by (rule power_mono) with \odd n\ show ?thesis by simp next - case False then have "0 \ b" by auto + case False + then have "0 \ b" by auto show ?thesis proof (cases "a < 0") - case True then have "n \ 0" and "a \ 0" using \odd n\ [THEN odd_pos] by auto + case True + then have "n \ 0" and "a \ 0" using \odd n\ [THEN odd_pos] by auto then have "a ^ n \ 0" unfolding power_le_zero_eq using \odd n\ by auto - moreover - from \0 \ b\ have "0 \ b ^ n" by auto + moreover from \0 \ b\ have "0 \ b ^ n" by auto ultimately show ?thesis by auto next - case False then have "0 \ a" by auto - with \a \ b\ show ?thesis using power_mono by auto + case False + then have "0 \ a" by auto + with \a \ b\ show ?thesis + using power_mono by auto qed qed @@ -347,13 +320,16 @@ by (fact zero_le_power_eq) lemma zero_less_power_eq_numeral [simp]: - "0 < a ^ numeral w \ numeral w = (0 :: nat) - \ even (numeral w :: nat) \ a \ 0 \ odd (numeral w :: nat) \ 0 < a" + "0 < a ^ numeral w \ + numeral w = (0 :: nat) \ + even (numeral w :: nat) \ a \ 0 \ + odd (numeral w :: nat) \ 0 < a" by (fact zero_less_power_eq) lemma power_le_zero_eq_numeral [simp]: - "a ^ numeral w \ 0 \ (0 :: nat) < numeral w - \ (odd (numeral w :: nat) \ a \ 0 \ even (numeral w :: nat) \ a = 0)" + "a ^ numeral w \ 0 \ + (0 :: nat) < numeral w \ + (odd (numeral w :: nat) \ a \ 0 \ even (numeral w :: nat) \ a = 0)" by (fact power_le_zero_eq) lemma power_less_zero_eq_numeral [simp]: @@ -367,10 +343,8 @@ end -subsubsection \Tools setup\ +subsubsection \Tool setup\ -declare transfer_morphism_int_nat [transfer add return: - even_int_iff -] +declare transfer_morphism_int_nat [transfer add return: even_int_iff] end diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Power.thy --- a/src/HOL/Power.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Power.thy Wed Aug 10 22:05:36 2016 +0200 @@ -6,7 +6,7 @@ section \Exponentiation\ theory Power -imports Num + imports Num begin subsection \Powers for Arbitrary Monoids\ @@ -15,9 +15,9 @@ begin primrec power :: "'a \ nat \ 'a" (infixr "^" 80) -where - power_0: "a ^ 0 = 1" -| power_Suc: "a ^ Suc n = a * a ^ n" + where + power_0: "a ^ 0 = 1" + | power_Suc: "a ^ Suc n = a * a ^ n" notation (latex output) power ("(_\<^bsup>_\<^esup>)" [1000] 1000) @@ -33,32 +33,25 @@ subclass power . -lemma power_one [simp]: - "1 ^ n = 1" +lemma power_one [simp]: "1 ^ n = 1" by (induct n) simp_all -lemma power_one_right [simp]: - "a ^ 1 = a" +lemma power_one_right [simp]: "a ^ 1 = a" by simp -lemma power_Suc0_right [simp]: - "a ^ Suc 0 = a" +lemma power_Suc0_right [simp]: "a ^ Suc 0 = a" by simp -lemma power_commutes: - "a ^ n * a = a * a ^ n" +lemma power_commutes: "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult.assoc) -lemma power_Suc2: - "a ^ Suc n = a ^ n * a" +lemma power_Suc2: "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) -lemma power_add: - "a ^ (m + n) = a ^ m * a ^ n" +lemma power_add: "a ^ (m + n) = a ^ m * a ^ n" by (induct m) (simp_all add: algebra_simps) -lemma power_mult: - "a ^ (m * n) = (a ^ m) ^ n" +lemma power_mult: "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) lemma power2_eq_square: "a\<^sup>2 = a * a" @@ -67,51 +60,49 @@ lemma power3_eq_cube: "a ^ 3 = a * a * a" by (simp add: numeral_3_eq_3 mult.assoc) -lemma power_even_eq: - "a ^ (2 * n) = (a ^ n)\<^sup>2" +lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2" by (subst mult.commute) (simp add: power_mult) -lemma power_odd_eq: - "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" +lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" by (simp add: power_even_eq) -lemma power_numeral_even: - "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" - unfolding numeral_Bit0 power_add Let_def .. +lemma power_numeral_even: "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)" + by (simp only: numeral_Bit0 power_add Let_def) -lemma power_numeral_odd: - "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" - unfolding numeral_Bit1 One_nat_def add_Suc_right add_0_right - unfolding power_Suc power_add Let_def mult.assoc .. +lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" + by (simp only: numeral_Bit1 One_nat_def add_Suc_right add_0_right + power_Suc power_add Let_def mult.assoc) -lemma funpow_times_power: - "(times x ^^ f x) = times (x ^ f x)" +lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)" proof (induct "f x" arbitrary: f) - case 0 then show ?case by (simp add: fun_eq_iff) + case 0 + then show ?case by (simp add: fun_eq_iff) next case (Suc n) define g where "g x = f x - 1" for x with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp - ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult.assoc) + ultimately show ?case + by (simp add: power_add funpow_add fun_eq_iff mult.assoc) qed lemma power_commuting_commutes: assumes "x * y = y * x" shows "x ^ n * y = y * x ^n" proof (induct n) + case 0 + then show ?case by simp +next case (Suc n) have "x ^ Suc n * y = x ^ n * y * x" by (subst power_Suc2) (simp add: assms ac_simps) also have "\ = y * x ^ Suc n" - unfolding Suc power_Suc2 - by (simp add: ac_simps) + by (simp only: Suc power_Suc2) (simp add: ac_simps) finally show ?case . -qed simp +qed -lemma power_minus_mult: - "0 < n \ a ^ (n - 1) * a = a ^ n" +lemma power_minus_mult: "0 < n \ a ^ (n - 1) * a = a ^ n" by (simp add: power_commutes split: nat_diff_split) end @@ -119,29 +110,25 @@ context comm_monoid_mult begin -lemma power_mult_distrib [field_simps]: - "(a * b) ^ n = (a ^ n) * (b ^ n)" +lemma power_mult_distrib [field_simps]: "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induct n) (simp_all add: ac_simps) end -text\Extract constant factors from powers\ +text \Extract constant factors from powers.\ declare power_mult_distrib [where a = "numeral w" for w, simp] declare power_mult_distrib [where b = "numeral w" for w, simp] -lemma power_add_numeral [simp]: - fixes a :: "'a :: monoid_mult" - shows "a^numeral m * a^numeral n = a^numeral (m + n)" +lemma power_add_numeral [simp]: "a^numeral m * a^numeral n = a^numeral (m + n)" + for a :: "'a::monoid_mult" by (simp add: power_add [symmetric]) -lemma power_add_numeral2 [simp]: - fixes a :: "'a :: monoid_mult" - shows "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b" +lemma power_add_numeral2 [simp]: "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b" + for a :: "'a::monoid_mult" by (simp add: mult.assoc [symmetric]) -lemma power_mult_numeral [simp]: - fixes a :: "'a :: monoid_mult" - shows"(a^numeral m)^numeral n = a^numeral (m * n)" +lemma power_mult_numeral [simp]: "(a^numeral m)^numeral n = a^numeral (m * n)" + for a :: "'a::monoid_mult" by (simp only: numeral_mult power_mult) context semiring_numeral @@ -151,8 +138,9 @@ by (simp only: sqr_conv_mult numeral_mult) lemma numeral_pow: "numeral (Num.pow k l) = numeral k ^ numeral l" - by (induct l, simp_all only: numeral_class.numeral.simps pow.simps - numeral_sqr numeral_mult power_add power_one_right) + by (induct l) + (simp_all only: numeral_class.numeral.simps pow.simps + numeral_sqr numeral_mult power_add power_one_right) lemma power_numeral [simp]: "numeral k ^ numeral l = numeral (Num.pow k l)" by (rule numeral_pow [symmetric]) @@ -162,16 +150,13 @@ context semiring_1 begin -lemma of_nat_power [simp]: - "of_nat (m ^ n) = of_nat m ^ n" +lemma of_nat_power [simp]: "of_nat (m ^ n) = of_nat m ^ n" by (induct n) simp_all -lemma zero_power: - "0 < n \ 0 ^ n = 0" +lemma zero_power: "0 < n \ 0 ^ n = 0" by (cases n) simp_all -lemma power_zero_numeral [simp]: - "0 ^ numeral k = 0" +lemma power_zero_numeral [simp]: "0 ^ numeral k = 0" by (simp add: numeral_eq_Suc) lemma zero_power2: "0\<^sup>2 = 0" (* delete? *) @@ -180,13 +165,11 @@ lemma one_power2: "1\<^sup>2 = 1" (* delete? *) by (rule power_one) -lemma power_0_Suc [simp]: - "0 ^ Suc n = 0" +lemma power_0_Suc [simp]: "0 ^ Suc n = 0" by simp -text\It looks plausible as a simprule, but its effect can be strange.\ -lemma power_0_left: - "0 ^ n = (if n = 0 then 1 else 0)" +text \It looks plausible as a simprule, but its effect can be strange.\ +lemma power_0_left: "0 ^ n = (if n = 0 then 1 else 0)" by (cases n) simp_all end @@ -194,34 +177,32 @@ context comm_semiring_1 begin -text \The divides relation\ +text \The divides relation.\ lemma le_imp_power_dvd: - assumes "m \ n" shows "a ^ m dvd a ^ n" + assumes "m \ n" + shows "a ^ m dvd a ^ n" proof - have "a ^ n = a ^ (m + (n - m))" - using \m \ n\ by simp - also have "\ = a ^ m * a ^ (n - m)" - by (rule power_add) + from assms have "a ^ n = a ^ (m + (n - m))" by simp + also have "\ = a ^ m * a ^ (n - m)" by (rule power_add) finally show "a ^ n = a ^ m * a ^ (n - m)" . qed -lemma power_le_dvd: - "a ^ n dvd b \ m \ n \ a ^ m dvd b" +lemma power_le_dvd: "a ^ n dvd b \ m \ n \ a ^ m dvd b" by (rule dvd_trans [OF le_imp_power_dvd]) -lemma dvd_power_same: - "x dvd y \ x ^ n dvd y ^ n" +lemma dvd_power_same: "x dvd y \ x ^ n dvd y ^ n" by (induct n) (auto simp add: mult_dvd_mono) -lemma dvd_power_le: - "x dvd y \ m \ n \ x ^ n dvd y ^ m" +lemma dvd_power_le: "x dvd y \ m \ n \ x ^ n dvd y ^ m" by (rule power_le_dvd [OF dvd_power_same]) lemma dvd_power [simp]: - assumes "n > (0::nat) \ x = 1" + fixes n :: nat + assumes "n > 0 \ x = 1" shows "x dvd (x ^ n)" -using assms proof + using assms +proof assume "0 < n" then have "x ^ n = x ^ Suc (n - 1)" by simp then show "x dvd (x ^ n)" by simp @@ -237,16 +218,13 @@ subclass power . -lemma power_eq_0_iff [simp]: - "a ^ n = 0 \ a = 0 \ n > 0" +lemma power_eq_0_iff [simp]: "a ^ n = 0 \ a = 0 \ n > 0" by (induct n) auto -lemma power_not_zero: - "a \ 0 \ a ^ n \ 0" +lemma power_not_zero: "a \ 0 \ a ^ n \ 0" by (induct n) auto -lemma zero_eq_power2 [simp]: - "a\<^sup>2 = 0 \ a = 0" +lemma zero_eq_power2 [simp]: "a\<^sup>2 = 0 \ a = 0" unfolding power2_eq_square by simp end @@ -254,45 +232,42 @@ context ring_1 begin -lemma power_minus: - "(- a) ^ n = (- 1) ^ n * a ^ n" +lemma power_minus: "(- a) ^ n = (- 1) ^ n * a ^ n" proof (induct n) - case 0 show ?case by simp + case 0 + show ?case by simp next - case (Suc n) then show ?case + case (Suc n) + then show ?case by (simp del: power_Suc add: power_Suc2 mult.assoc) qed lemma power_minus': "NO_MATCH 1 x \ (-x) ^ n = (-1)^n * x ^ n" by (rule power_minus) -lemma power_minus_Bit0: - "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)" +lemma power_minus_Bit0: "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)" by (induct k, simp_all only: numeral_class.numeral.simps power_add power_one_right mult_minus_left mult_minus_right minus_minus) -lemma power_minus_Bit1: - "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))" +lemma power_minus_Bit1: "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))" by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left) -lemma power2_minus [simp]: - "(- a)\<^sup>2 = a\<^sup>2" +lemma power2_minus [simp]: "(- a)\<^sup>2 = a\<^sup>2" by (fact power_minus_Bit0) -lemma power_minus1_even [simp]: - "(- 1) ^ (2*n) = 1" +lemma power_minus1_even [simp]: "(- 1) ^ (2*n) = 1" proof (induct n) - case 0 show ?case by simp + case 0 + show ?case by simp next - case (Suc n) then show ?case by (simp add: power_add power2_eq_square) + case (Suc n) + then show ?case by (simp add: power_add power2_eq_square) qed -lemma power_minus1_odd: - "(- 1) ^ Suc (2*n) = -1" +lemma power_minus1_odd: "(- 1) ^ Suc (2*n) = -1" by simp -lemma power_minus_even [simp]: - "(-a) ^ (2*n) = a ^ (2*n)" +lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)" by (simp add: power_minus [of a]) end @@ -300,8 +275,7 @@ context ring_1_no_zero_divisors begin -lemma power2_eq_1_iff: - "a\<^sup>2 = 1 \ a = 1 \ a = - 1" +lemma power2_eq_1_iff: "a\<^sup>2 = 1 \ a = 1 \ a = - 1" using square_eq_1_iff [of a] by (simp add: power2_eq_square) end @@ -317,13 +291,10 @@ context algebraic_semidom begin -lemma div_power: - assumes "b dvd a" - shows "(a div b) ^ n = a ^ n div b ^ n" - using assms by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same) +lemma div_power: "b dvd a \ (a div b) ^ n = a ^ n div b ^ n" + by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same) -lemma is_unit_power_iff: - "is_unit (a ^ n) \ is_unit a \ n = 0" +lemma is_unit_power_iff: "is_unit (a ^ n) \ is_unit a \ n = 0" by (induct n) (auto simp add: is_unit_mult_iff) end @@ -331,12 +302,10 @@ context normalization_semidom begin -lemma normalize_power: - "normalize (a ^ n) = normalize a ^ n" +lemma normalize_power: "normalize (a ^ n) = normalize a ^ n" by (induct n) (simp_all add: normalize_mult) -lemma unit_factor_power: - "unit_factor (a ^ n) = unit_factor a ^ n" +lemma unit_factor_power: "unit_factor (a ^ n) = unit_factor a ^ n" by (induct n) (simp_all add: unit_factor_mult) end @@ -344,19 +313,19 @@ context division_ring begin -text\Perhaps these should be simprules.\ -lemma power_inverse [field_simps, divide_simps]: - "inverse a ^ n = inverse (a ^ n)" +text \Perhaps these should be simprules.\ +lemma power_inverse [field_simps, divide_simps]: "inverse a ^ n = inverse (a ^ n)" proof (cases "a = 0") - case True then show ?thesis by (simp add: power_0_left) + case True + then show ?thesis by (simp add: power_0_left) next - case False then have "inverse (a ^ n) = inverse a ^ n" + case False + then have "inverse (a ^ n) = inverse a ^ n" by (induct n) (simp_all add: nonzero_inverse_mult_distrib power_commutes) then show ?thesis by simp qed -lemma power_one_over [field_simps, divide_simps]: - "(1 / a) ^ n = 1 / a ^ n" +lemma power_one_over [field_simps, divide_simps]: "(1 / a) ^ n = 1 / a ^ n" using power_inverse [of a] by (simp add: divide_inverse) end @@ -365,12 +334,11 @@ begin lemma power_diff: - assumes nz: "a \ 0" + assumes "a \ 0" shows "n \ m \ a ^ (m - n) = a ^ m / a ^ n" - by (induct m n rule: diff_induct) (simp_all add: nz power_not_zero) + by (induct m n rule: diff_induct) (simp_all add: assms power_not_zero) -lemma power_divide [field_simps, divide_simps]: - "(a / b) ^ n = a ^ n / b ^ n" +lemma power_divide [field_simps, divide_simps]: "(a / b) ^ n = a ^ n / b ^ n" by (induct n) simp_all end @@ -381,22 +349,19 @@ context linordered_semidom begin -lemma zero_less_power [simp]: - "0 < a \ 0 < a ^ n" +lemma zero_less_power [simp]: "0 < a \ 0 < a ^ n" by (induct n) simp_all -lemma zero_le_power [simp]: - "0 \ a \ 0 \ a ^ n" +lemma zero_le_power [simp]: "0 \ a \ 0 \ a ^ n" by (induct n) simp_all -lemma power_mono: - "a \ b \ 0 \ a \ a ^ n \ b ^ n" +lemma power_mono: "a \ b \ 0 \ a \ a ^ n \ b ^ n" by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" using power_mono [of 1 a n] by simp -lemma power_le_one: "\0 \ a; a \ 1\ \ a ^ n \ 1" +lemma power_le_one: "0 \ a \ a \ 1 \ a ^ n \ 1" using power_mono [of a 1 n] by simp lemma power_gt1_lemma: @@ -405,19 +370,16 @@ proof - from gt1 have "0 \ a" by (fact order_trans [OF zero_le_one less_imp_le]) - have "1 * 1 < a * 1" using gt1 by simp - also have "\ \ a * a ^ n" using gt1 - by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le - zero_le_one order_refl) + from gt1 have "1 * 1 < a * 1" by simp + also from gt1 have "\ \ a * a ^ n" + by (simp only: mult_mono \0 \ a\ one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed -lemma power_gt1: - "1 < a \ 1 < a ^ Suc n" +lemma power_gt1: "1 < a \ 1 < a ^ Suc n" by (simp add: power_gt1_lemma) -lemma one_less_power [simp]: - "1 < a \ 0 < n \ 1 < a ^ n" +lemma one_less_power [simp]: "1 < a \ 0 < n \ 1 < a ^ n" by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: @@ -431,123 +393,122 @@ show ?case proof (cases n) case 0 - with Suc.prems Suc.hyps have "a * a ^ m \ 1" by simp + with Suc have "a * a ^ m \ 1" by simp with gt1 show ?thesis - by (force simp only: power_gt1_lemma - not_less [symmetric]) + by (force simp only: power_gt1_lemma not_less [symmetric]) next case (Suc n) with Suc.prems Suc.hyps show ?thesis - by (force dest: mult_left_le_imp_le - simp add: less_trans [OF zero_less_one gt1]) + by (force dest: mult_left_le_imp_le simp add: less_trans [OF zero_less_one gt1]) qed qed -lemma of_nat_zero_less_power_iff [simp]: - "of_nat x ^ n > 0 \ x > 0 \ n = 0" +lemma of_nat_zero_less_power_iff [simp]: "of_nat x ^ n > 0 \ x > 0 \ n = 0" by (induct n) auto -text\Surely we can strengthen this? It holds for \0 too.\ -lemma power_inject_exp [simp]: - "1 < a \ a ^ m = a ^ n \ m = n" +text \Surely we can strengthen this? It holds for \0 too.\ +lemma power_inject_exp [simp]: "1 < a \ a ^ m = a ^ n \ m = n" by (force simp add: order_antisym power_le_imp_le_exp) -text\Can relax the first premise to @{term "0 -lemma power_less_imp_less_exp: - "1 < a \ a ^ m < a ^ n \ m < n" - by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] - power_le_imp_le_exp) +text \ + Can relax the first premise to @{term "0 +lemma power_less_imp_less_exp: "1 < a \ a ^ m < a ^ n \ m < n" + by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] power_le_imp_le_exp) -lemma power_strict_mono [rule_format]: - "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" - by (induct n) - (auto simp add: mult_strict_mono le_less_trans [of 0 a b]) +lemma power_strict_mono [rule_format]: "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" + by (induct n) (auto simp: mult_strict_mono le_less_trans [of 0 a b]) text\Lemma for \power_strict_decreasing\\ -lemma power_Suc_less: - "0 < a \ a < 1 \ a * a ^ n < a ^ n" - by (induct n) - (auto simp add: mult_strict_left_mono) +lemma power_Suc_less: "0 < a \ a < 1 \ a * a ^ n < a ^ n" + by (induct n) (auto simp: mult_strict_left_mono) -lemma power_strict_decreasing [rule_format]: - "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" +lemma power_strict_decreasing [rule_format]: "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" proof (induct N) - case 0 then show ?case by simp + case 0 + then show ?case by simp next - case (Suc N) then show ?case - apply (auto simp add: power_Suc_less less_Suc_eq) - apply (subgoal_tac "a * a^N < 1 * a^n") - apply simp - apply (rule mult_strict_mono) apply auto - done + case (Suc N) + then show ?case + apply (auto simp add: power_Suc_less less_Suc_eq) + apply (subgoal_tac "a * a^N < 1 * a^n") + apply simp + apply (rule mult_strict_mono) + apply auto + done qed -text\Proof resembles that of \power_strict_decreasing\\ -lemma power_decreasing [rule_format]: - "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" +text \Proof resembles that of \power_strict_decreasing\.\ +lemma power_decreasing: "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" proof (induct N) - case 0 then show ?case by simp + case 0 + then show ?case by simp next - case (Suc N) then show ?case - apply (auto simp add: le_Suc_eq) - apply (subgoal_tac "a * a^N \ 1 * a^n", simp) - apply (rule mult_mono) apply auto - done + case (Suc N) + then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "a * a^N \ 1 * a^n") + apply simp + apply (rule mult_mono) + apply auto + done qed -lemma power_Suc_less_one: - "0 < a \ a < 1 \ a ^ Suc n < 1" +lemma power_Suc_less_one: "0 < a \ a < 1 \ a ^ Suc n < 1" using power_strict_decreasing [of 0 "Suc n" a] by simp -text\Proof again resembles that of \power_strict_decreasing\\ -lemma power_increasing [rule_format]: - "n \ N \ 1 \ a \ a ^ n \ a ^ N" +text \Proof again resembles that of \power_strict_decreasing\.\ +lemma power_increasing: "n \ N \ 1 \ a \ a ^ n \ a ^ N" proof (induct N) - case 0 then show ?case by simp + case 0 + then show ?case by simp next - case (Suc N) then show ?case - apply (auto simp add: le_Suc_eq) - apply (subgoal_tac "1 * a^n \ a * a^N", simp) - apply (rule mult_mono) apply (auto simp add: order_trans [OF zero_le_one]) - done + case (Suc N) + then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "1 * a^n \ a * a^N") + apply simp + apply (rule mult_mono) + apply (auto simp add: order_trans [OF zero_le_one]) + done qed -text\Lemma for \power_strict_increasing\\ -lemma power_less_power_Suc: - "1 < a \ a ^ n < a * a ^ n" - by (induct n) (auto simp add: mult_strict_left_mono less_trans [OF zero_less_one]) +text \Lemma for \power_strict_increasing\.\ +lemma power_less_power_Suc: "1 < a \ a ^ n < a * a ^ n" + by (induct n) (auto simp: mult_strict_left_mono less_trans [OF zero_less_one]) -lemma power_strict_increasing [rule_format]: - "n < N \ 1 < a \ a ^ n < a ^ N" +lemma power_strict_increasing: "n < N \ 1 < a \ a ^ n < a ^ N" proof (induct N) - case 0 then show ?case by simp + case 0 + then show ?case by simp next - case (Suc N) then show ?case - apply (auto simp add: power_less_power_Suc less_Suc_eq) - apply (subgoal_tac "1 * a^n < a * a^N", simp) - apply (rule mult_strict_mono) apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) - done + case (Suc N) + then show ?case + apply (auto simp add: power_less_power_Suc less_Suc_eq) + apply (subgoal_tac "1 * a^n < a * a^N") + apply simp + apply (rule mult_strict_mono) + apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) + done qed -lemma power_increasing_iff [simp]: - "1 < b \ b ^ x \ b ^ y \ x \ y" +lemma power_increasing_iff [simp]: "1 < b \ b ^ x \ b ^ y \ x \ y" by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) -lemma power_strict_increasing_iff [simp]: - "1 < b \ b ^ x < b ^ y \ x < y" -by (blast intro: power_less_imp_less_exp power_strict_increasing) +lemma power_strict_increasing_iff [simp]: "1 < b \ b ^ x < b ^ y \ x < y" + by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: assumes le: "a ^ Suc n \ b ^ Suc n" - and ynonneg: "0 \ b" + and "0 \ b" shows "a \ b" proof (rule ccontr) - assume "~ a \ b" + assume "\ ?thesis" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" - by (simp only: assms power_strict_mono) - from le and this show False + by (simp only: assms(2) power_strict_mono) + with le show False by (simp add: linorder_not_less [symmetric]) qed @@ -556,38 +517,31 @@ assumes nonneg: "0 \ b" shows "a < b" proof (rule contrapos_pp [OF less]) - assume "~ a < b" - hence "b \ a" by (simp only: linorder_not_less) - hence "b ^ n \ a ^ n" using nonneg by (rule power_mono) - thus "\ a ^ n < b ^ n" by (simp only: linorder_not_less) + assume "\ ?thesis" + then have "b \ a" by (simp only: linorder_not_less) + from this nonneg have "b ^ n \ a ^ n" by (rule power_mono) + then show "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed -lemma power_inject_base: - "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" -by (blast intro: power_le_imp_le_base antisym eq_refl sym) +lemma power_inject_base: "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" + by (blast intro: power_le_imp_le_base antisym eq_refl sym) -lemma power_eq_imp_eq_base: - "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" +lemma power_eq_imp_eq_base: "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" by (cases n) (simp_all del: power_Suc, rule power_inject_base) -lemma power_eq_iff_eq_base: - "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" +lemma power_eq_iff_eq_base: "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" using power_eq_imp_eq_base [of a n b] by auto -lemma power2_le_imp_le: - "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" +lemma power2_le_imp_le: "x\<^sup>2 \ y\<^sup>2 \ 0 \ y \ x \ y" unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) -lemma power2_less_imp_less: - "x\<^sup>2 < y\<^sup>2 \ 0 \ y \ x < y" +lemma power2_less_imp_less: "x\<^sup>2 < y\<^sup>2 \ 0 \ y \ x < y" by (rule power_less_imp_less_base) -lemma power2_eq_imp_eq: - "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" +lemma power2_eq_imp_eq: "x\<^sup>2 = y\<^sup>2 \ 0 \ x \ 0 \ y \ x = y" unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp -lemma power_Suc_le_self: - shows "0 \ a \ a \ 1 \ a ^ Suc n \ a" +lemma power_Suc_le_self: "0 \ a \ a \ 1 \ a ^ Suc n \ a" using power_decreasing [of 1 "Suc n" a] by simp end @@ -595,16 +549,13 @@ context linordered_ring_strict begin -lemma sum_squares_eq_zero_iff: - "x * x + y * y = 0 \ x = 0 \ y = 0" +lemma sum_squares_eq_zero_iff: "x * x + y * y = 0 \ x = 0 \ y = 0" by (simp add: add_nonneg_eq_0_iff) -lemma sum_squares_le_zero_iff: - "x * x + y * y \ 0 \ x = 0 \ y = 0" +lemma sum_squares_le_zero_iff: "x * x + y * y \ 0 \ x = 0 \ y = 0" by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) -lemma sum_squares_gt_zero_iff: - "0 < x * x + y * y \ x \ 0 \ y \ 0" +lemma sum_squares_gt_zero_iff: "0 < x * x + y * y \ x \ 0 \ y \ 0" by (simp add: not_le [symmetric] sum_squares_le_zero_iff) end @@ -620,28 +571,26 @@ lemma zero_less_power_abs_iff [simp]: "0 < \a\ ^ n \ a \ 0 \ n = 0" proof (induct n) - case 0 show ?case by simp + case 0 + show ?case by simp next - case (Suc n) show ?case by (auto simp add: Suc zero_less_mult_iff) + case Suc + then show ?case by (auto simp: zero_less_mult_iff) qed lemma zero_le_power_abs [simp]: "0 \ \a\ ^ n" by (rule zero_le_power [OF abs_ge_zero]) -lemma zero_le_power2 [simp]: - "0 \ a\<^sup>2" +lemma zero_le_power2 [simp]: "0 \ a\<^sup>2" by (simp add: power2_eq_square) -lemma zero_less_power2 [simp]: - "0 < a\<^sup>2 \ a \ 0" +lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \ a \ 0" by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) -lemma power2_less_0 [simp]: - "\ a\<^sup>2 < 0" +lemma power2_less_0 [simp]: "\ a\<^sup>2 < 0" by (force simp add: power2_eq_square mult_less_0_iff) -lemma power2_less_eq_zero_iff [simp]: - "a\<^sup>2 \ 0 \ a = 0" +lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \ 0 \ a = 0" by (simp add: le_less) lemma abs_power2 [simp]: "\a\<^sup>2\ = a\<^sup>2" @@ -650,8 +599,7 @@ lemma power2_abs [simp]: "\a\\<^sup>2 = a\<^sup>2" by (simp add: power2_eq_square) -lemma odd_power_less_zero: - "a < 0 \ a ^ Suc (2*n) < 0" +lemma odd_power_less_zero: "a < 0 \ a ^ Suc (2*n) < 0" proof (induct n) case 0 then show ?case by simp @@ -659,160 +607,152 @@ case (Suc n) have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" by (simp add: ac_simps power_add power2_eq_square) - thus ?case + then show ?case by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg) qed -lemma odd_0_le_power_imp_0_le: - "0 \ a ^ Suc (2*n) \ 0 \ a" +lemma odd_0_le_power_imp_0_le: "0 \ a ^ Suc (2*n) \ 0 \ a" using odd_power_less_zero [of a n] - by (force simp add: linorder_not_less [symmetric]) + by (force simp add: linorder_not_less [symmetric]) -lemma zero_le_even_power'[simp]: - "0 \ a ^ (2*n)" +lemma zero_le_even_power'[simp]: "0 \ a ^ (2*n)" proof (induct n) case 0 - show ?case by simp + show ?case by simp next case (Suc n) - have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" - by (simp add: ac_simps power_add power2_eq_square) - thus ?case - by (simp add: Suc zero_le_mult_iff) + have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" + by (simp add: ac_simps power_add power2_eq_square) + then show ?case + by (simp add: Suc zero_le_mult_iff) qed -lemma sum_power2_ge_zero: - "0 \ x\<^sup>2 + y\<^sup>2" +lemma sum_power2_ge_zero: "0 \ x\<^sup>2 + y\<^sup>2" by (intro add_nonneg_nonneg zero_le_power2) -lemma not_sum_power2_lt_zero: - "\ x\<^sup>2 + y\<^sup>2 < 0" +lemma not_sum_power2_lt_zero: "\ x\<^sup>2 + y\<^sup>2 < 0" unfolding not_less by (rule sum_power2_ge_zero) -lemma sum_power2_eq_zero_iff: - "x\<^sup>2 + y\<^sup>2 = 0 \ x = 0 \ y = 0" +lemma sum_power2_eq_zero_iff: "x\<^sup>2 + y\<^sup>2 = 0 \ x = 0 \ y = 0" unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff) -lemma sum_power2_le_zero_iff: - "x\<^sup>2 + y\<^sup>2 \ 0 \ x = 0 \ y = 0" +lemma sum_power2_le_zero_iff: "x\<^sup>2 + y\<^sup>2 \ 0 \ x = 0 \ y = 0" by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero) -lemma sum_power2_gt_zero_iff: - "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" +lemma sum_power2_gt_zero_iff: "0 < x\<^sup>2 + y\<^sup>2 \ x \ 0 \ y \ 0" unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff) -lemma abs_le_square_iff: - "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" +lemma abs_le_square_iff: "\x\ \ \y\ \ x\<^sup>2 \ y\<^sup>2" + (is "?lhs \ ?rhs") proof - assume "\x\ \ \y\" - then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono, simp) - then show "x\<^sup>2 \ y\<^sup>2" by simp + assume ?lhs + then have "\x\\<^sup>2 \ \y\\<^sup>2" by (rule power_mono) simp + then show ?rhs by simp next - assume "x\<^sup>2 \ y\<^sup>2" - then show "\x\ \ \y\" + assume ?rhs + then show ?lhs by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero]) qed lemma abs_square_le_1:"x\<^sup>2 \ 1 \ \x\ \ 1" - using abs_le_square_iff [of x 1] - by simp + using abs_le_square_iff [of x 1] by simp lemma abs_square_eq_1: "x\<^sup>2 = 1 \ \x\ = 1" by (auto simp add: abs_if power2_eq_1_iff) lemma abs_square_less_1: "x\<^sup>2 < 1 \ \x\ < 1" - using abs_square_eq_1 [of x] abs_square_le_1 [of x] - by (auto simp add: le_less) + using abs_square_eq_1 [of x] abs_square_le_1 [of x] by (auto simp add: le_less) end subsection \Miscellaneous rules\ -lemma (in linordered_semidom) self_le_power: - "1 \ a \ 0 < n \ a \ a ^ n" +lemma (in linordered_semidom) self_le_power: "1 \ a \ 0 < n \ a \ a ^ n" using power_increasing [of 1 n a] power_one_right [of a] by auto -lemma (in power) power_eq_if: - "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))" +lemma (in power) power_eq_if: "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))" unfolding One_nat_def by (cases m) simp_all -lemma (in comm_semiring_1) power2_sum: - "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y" +lemma (in comm_semiring_1) power2_sum: "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) -lemma (in comm_ring_1) power2_diff: - "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" +context comm_ring_1 +begin + +lemma power2_diff: "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square mult_2_right) -lemma (in comm_ring_1) power2_commute: - "(x - y)\<^sup>2 = (y - x)\<^sup>2" +lemma power2_commute: "(x - y)\<^sup>2 = (y - x)\<^sup>2" by (simp add: algebra_simps power2_eq_square) -lemma (in comm_ring_1) minus_power_mult_self: - "(- a) ^ n * (- a) ^ n = a ^ (2 * n)" - by (simp add: power_mult_distrib [symmetric]) (simp add: power2_eq_square [symmetric] power_mult [symmetric]) - -lemma (in comm_ring_1) minus_one_mult_self [simp]: - "(- 1) ^ n * (- 1) ^ n = 1" +lemma minus_power_mult_self: "(- a) ^ n * (- a) ^ n = a ^ (2 * n)" + by (simp add: power_mult_distrib [symmetric]) + (simp add: power2_eq_square [symmetric] power_mult [symmetric]) + +lemma minus_one_mult_self [simp]: "(- 1) ^ n * (- 1) ^ n = 1" using minus_power_mult_self [of 1 n] by simp -lemma (in comm_ring_1) left_minus_one_mult_self [simp]: - "(- 1) ^ n * ((- 1) ^ n * a) = a" +lemma left_minus_one_mult_self [simp]: "(- 1) ^ n * ((- 1) ^ n * a) = a" by (simp add: mult.assoc [symmetric]) +end + text \Simprules for comparisons where common factors can be cancelled.\ lemmas zero_compare_simps = - add_strict_increasing add_strict_increasing2 add_increasing - zero_le_mult_iff zero_le_divide_iff - zero_less_mult_iff zero_less_divide_iff - mult_le_0_iff divide_le_0_iff - mult_less_0_iff divide_less_0_iff - zero_le_power2 power2_less_0 + add_strict_increasing add_strict_increasing2 add_increasing + zero_le_mult_iff zero_le_divide_iff + zero_less_mult_iff zero_less_divide_iff + mult_le_0_iff divide_le_0_iff + mult_less_0_iff divide_less_0_iff + zero_le_power2 power2_less_0 subsection \Exponentiation for the Natural Numbers\ -lemma nat_one_le_power [simp]: - "Suc 0 \ i \ Suc 0 \ i ^ n" +lemma nat_one_le_power [simp]: "Suc 0 \ i \ Suc 0 \ i ^ n" by (rule one_le_power [of i n, unfolded One_nat_def]) -lemma nat_zero_less_power_iff [simp]: - "x ^ n > 0 \ x > (0::nat) \ n = 0" +lemma nat_zero_less_power_iff [simp]: "x ^ n > 0 \ x > 0 \ n = 0" + for x :: nat by (induct n) auto -lemma nat_power_eq_Suc_0_iff [simp]: - "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" +lemma nat_power_eq_Suc_0_iff [simp]: "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" by (induct m) auto -lemma power_Suc_0 [simp]: - "Suc 0 ^ n = Suc 0" +lemma power_Suc_0 [simp]: "Suc 0 ^ n = Suc 0" by simp -text\Valid for the naturals, but what if \0? -Premises cannot be weakened: consider the case where @{term "i=0"}, -@{term "m=1"} and @{term "n=0"}.\ +text \ + Valid for the naturals, but what if \0 < i < 1\? Premises cannot be + weakened: consider the case where \i = 0\, \m = 1\ and \n = 0\. +\ + lemma nat_power_less_imp_less: - assumes nonneg: "0 < (i::nat)" + fixes i :: nat + assumes nonneg: "0 < i" assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") - case True with less power_one [where 'a = nat] show ?thesis by simp + case True + with less power_one [where 'a = nat] show ?thesis by simp next - case False with nonneg have "1 < i" by auto + case False + with nonneg have "1 < i" by auto from power_strict_increasing_iff [OF this] less show ?thesis .. qed -lemma power_dvd_imp_le: - "i ^ m dvd i ^ n \ (1::nat) < i \ m \ n" - apply (rule power_le_imp_le_exp, assumption) - apply (erule dvd_imp_le, simp) +lemma power_dvd_imp_le: "i ^ m dvd i ^ n \ 1 < i \ m \ n" + for i m n :: nat + apply (rule power_le_imp_le_exp) + apply assumption + apply (erule dvd_imp_le) + apply simp done -lemma power2_nat_le_eq_le: - fixes m n :: nat - shows "m\<^sup>2 \ n\<^sup>2 \ m \ n" +lemma power2_nat_le_eq_le: "m\<^sup>2 \ n\<^sup>2 \ m \ n" + for m n :: nat by (auto intro: power2_le_imp_le power_mono) lemma power2_nat_le_imp_le: @@ -820,18 +760,20 @@ assumes "m\<^sup>2 \ n" shows "m \ n" proof (cases m) - case 0 then show ?thesis by simp + case 0 + then show ?thesis by simp next case (Suc k) show ?thesis proof (rule ccontr) - assume "\ m \ n" + assume "\ ?thesis" then have "n < m" by simp with assms Suc show False by (simp add: power2_eq_square) qed qed + subsubsection \Cardinality of the Powerset\ lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2" @@ -840,16 +782,17 @@ lemma card_Pow: "finite A \ card (Pow A) = 2 ^ card A" proof (induct rule: finite_induct) case empty - show ?case by auto + show ?case by auto next case (insert x A) then have "inj_on (insert x) (Pow A)" unfolding inj_on_def by (blast elim!: equalityE) then have "card (Pow A) + card (insert x ` Pow A) = 2 * 2 ^ card A" by (simp add: mult_2 card_image Pow_insert insert.hyps) - then show ?case using insert + with insert show ?case apply (simp add: Pow_insert) - apply (subst card_Un_disjoint, auto) + apply (subst card_Un_disjoint) + apply auto done qed