(* Title: HOL/Set_Interval.thy Author: Tobias Nipkow Author: Clemens Ballarin Author: Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals *) header {* Set intervals *} theory Set_Interval imports Int Nat_Transfer begin context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text{* A note of warning when using @{term"{.. 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (xsymbols) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\ _\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\ _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\ _\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\ _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) translations "UN i<=n. A" == "UN i:{..n}. A" "UN i atLeast y) = (y \ (x::'a::order))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::linorder))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" apply (auto simp add: greaterThan_def) apply (subst linorder_not_less [symmetric], blast) done lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" apply (rule iffI) apply (erule equalityE) apply simp_all done lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::order))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" apply (auto simp add: lessThan_def) apply (subst linorder_not_less [symmetric], blast) done lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" apply (rule iffI) apply (erule equalityE) apply simp_all done lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) subsection {*Two-sided intervals*} context ord begin lemma greaterThanLessThan_iff [simp,no_atp]: "(i : {l<.. {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (~ a <= b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (~ a <= b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty[simp]: "b <= a \ {a.. (~ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (~ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty[simp]: "l \ k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ ~ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ ~ k < l" by auto (blast intro: less_le_trans) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma atLeastatMost_subset_iff[simp]: "{a..b} <= {c..d} \ (~ a <= b) | c <= a & b <= d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((~ a <= b) | c <= a & b <= d & (c < a | b < d)) & c <= d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_singleton_iff[simp]: "{a .. b} = {c} \ a = b \ b = c" proof assume "{a..b} = {c}" hence "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp moreover with `{a..b} = {c}` have "c \ a \ b \ c" by auto ultimately show "a = b \ b = c" by auto qed simp end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a .. b} \ { c ..< d } \ (a \ b \ c \ a \ b < d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. b <= a | c<=a & b<=d" apply (auto simp:subset_eq Ball_def) apply(frule_tac x=a in spec) apply(erule_tac x=d in allE) apply (simp add: less_imp_le) done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le linorder_antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto subsubsection {* Intersection *} context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) end subsection {* Intervals of natural numbers *} subsubsection {* The Constant @{term lessThan} *} lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text {* The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform @{term "{.. Suc ` {.. Suc (x - 1)" by auto with `x < Suc n` show "x = 0" by auto qed lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" by (simp add: lessThan_def atMost_def less_Suc_eq_le) lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV" by blast subsubsection {* The Constant @{term greaterThan} *} lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc" apply (simp add: greaterThan_def) apply (blast dest: gr0_conv_Suc [THEN iffD1]) done lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" apply (simp add: greaterThan_def) apply (auto elim: linorder_neqE) done lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}" by blast subsubsection {* The Constant @{term atLeast} *} lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" apply (simp add: atLeast_def) apply (simp add: Suc_le_eq) apply (simp add: order_le_less, blast) done lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV" by blast subsubsection {* The Constant @{term atMost} *} lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" apply (simp add: atMost_def) apply (simp add: less_Suc_eq order_le_less, blast) done lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV" by blast subsubsection {* The Constant @{term atLeastLessThan} *} text{*The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one. *} lemma atLeast0LessThan: "{0::nat.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by (auto simp add: atLeastAtMost_def) lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text {* The analogous result is useful on @{typ int}: *} (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m <= 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j.. ?B" by auto next show "?B \ ?A" proof fix n assume a: "n : ?B" hence "n - k : {i..j}" by auto moreover have "n = (n - k) + k" using a by auto ultimately show "n : ?A" by blast qed qed lemma image_add_atLeastLessThan: "(%n::nat. n+k) ` {i.. ?B" by auto next show "?B \ ?A" proof fix n assume a: "n : ?B" hence "n - k : {i..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: split_if_asm) qed qed auto context ordered_ab_group_add begin lemma fixes x :: 'a shows image_uminus_greaterThan[simp]: "uminus ` {x<..} = {..<-x}" and image_uminus_atLeast[simp]: "uminus ` {x..} = {..-x}" proof safe fix y assume "y < -x" hence *: "x < -y" using neg_less_iff_less[of "-y" x] by simp have "- (-y) \ uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (insert `y \ -x`[THEN le_imp_neg_le], simp) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. finite N" apply (rule finite_subset) apply (rule_tac [2] finite_lessThan, auto) done text {* A set of natural numbers is finite iff it is bounded. *} lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (EX m. ALL n:N. nnat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) text{* Any subset of an interval of natural numbers the size of the subset is exactly that interval. *} lemma subset_card_intvl_is_intvl: "A <= {k.. A = {k..i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A <= ?B" proof fix x assume "x : ?A" then obtain i where i: "i\n" "x : M i" by auto show "x : ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B <= ?A" by auto qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A <= ?B" by fastforce next show "?B <= ?A" proof fix x assume "x : ?B" then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto hence "i-k\n & x : M((i-k)+k)" by auto thus "x : ?A" by blast qed qed lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(!!n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: "(!!n::nat. (\i\{0.. (\i\{0.. (\n. A n) \ (\n. B n)" apply (rule UN_finite_subset) apply (subst UN_UN_finite_eq [symmetric, of B]) apply blast done lemma UN_finite2_eq: "(!!n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" apply (rule subset_antisym) apply (rule UN_finite2_subset, blast) apply (rule UN_finite2_subset [where k=k]) apply (force simp add: atLeastLessThan_add_Un [of 0]) done subsubsection {* Cardinality *} lemma card_lessThan [simp]: "card {.. \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ EX h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes FIN: "finite A" and FIN': "finite B" shows BIJ: "(\f. bij_betw f A B) \ (card A = card B)" using assms proof(auto simp add: bij_betw_same_card) assume *: "card A = card B" obtain f where "bij_betw f A {0 ..< card A}" using FIN ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using FIN' ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g o f) A B" using * by (auto simp add: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed lemma inj_on_iff_card_le: assumes FIN: "finite A" and FIN': "finite B" shows "(\f. inj_on f A \ f ` A \ B) = (card A \ card B)" proof (safe intro!: card_inj_on_le) assume *: "card A \ card B" obtain f where 1: "inj_on f A" and 2: "f ` A = {0 ..< card A}" using FIN ex_bij_betw_finite_nat unfolding bij_betw_def by force moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B" using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force hence "inj_on (g o f) A" using 1 comp_inj_on by blast moreover {have "{0 ..< card A} \ {0 ..< card B}" using * by force with 2 have "f ` A \ {0 ..< card B}" by blast hence "(g o f) ` A \ B" unfolding comp_def using 3 by force } ultimately show "(\f. inj_on f A \ f ` A \ B)" by blast qed (insert assms, auto) subsection {* Intervals of integers *} lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l.. u ==> {(0::int).. u") apply (subst image_atLeastZeroLessThan_int, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_int [iff]: "finite {l.. u") apply (subst image_atLeastZeroLessThan_int, assumption) apply (subst card_image) apply (auto simp add: inj_on_def) done lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: "0 \ M \ card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" apply (rule card_bij_eq [of Suc _ _ "\x. x - Suc 0"]) apply simp apply fastforce apply auto apply (rule inj_on_diff_nat) apply auto apply (case_tac x) apply auto apply (case_tac xa) apply auto apply (case_tac xa) apply auto done lemma card_less_Suc: assumes zero_in_M: "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - from assms have a: "0 \ {k \ M. k < Suc i}" by simp hence c: "{k \ M. k < Suc i} = insert 0 ({k \ M. k < Suc i} - {0})" by (auto simp only: insert_Diff) have b: "{k \ M. k < Suc i} - {0} = {k \ M - {0}. k < Suc i}" by auto from finite_M_bounded_by_nat[of "\x. x \ M" "Suc i"] have "Suc (card {k. Suc k \ M \ k < i}) = card (insert 0 ({k \ M. k < Suc i} - {0}))" apply (subst card_insert) apply simp_all apply (subst b) apply (subst card_less_Suc2[symmetric]) apply simp_all done with c show ?thesis by simp qed subsection {*Lemmas useful with the summation operator setsum*} text {* For examples, see Algebra/poly/UnivPoly2.thy *} subsubsection {* Disjoint Unions *} text {* Singletons and open intervals *} lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. {l} Un {l<..u} = {l..u}" "(l::'a::linorder) <= u ==> {l.. {..l} Un {l<.. {.. {..l} Un {l<..u} = {..u}" "(l::'a::linorder) <= u ==> {.. {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. {l..u} Un {u<..} = {l..}" "(l::'a::linorder) <= u ==> {l.. {l<.. {l<..m} Un {m<.. {l.. {l..m} Un {m<.. {l<.. {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) <= m; m <= u |] ==> {l.. {l..m} Un {m<..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two subsubsection {* Disjoint Intersections *} text {* One- and two-sided intervals *} lemma ivl_disj_int_one: "{..l::'a::order} Int {l<.. n \ {i.. {m.. i | m \ i & j \ (n::'a::linorder))" apply(auto simp:linorder_not_le) apply(rule ccontr) apply(insert linorder_le_less_linear[of i n]) apply(clarsimp simp:linorder_not_le) apply(fastforce) done subsection {* Summation indexed over intervals *} syntax "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (xsymbols) "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) syntax (HTML output) "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10) "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10) "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\sum_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST setsum (%x. t) {a..b}" "\x=a..i\n. t" == "CONST setsum (\i. t) {..n}" "\ii. t) {..x\{a..b}. e"} & @{term"\x=a..b. e"} & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..x=a..x=a..x\{..b}. e"} & @{term"\x\b. e"} & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..xxx::nat=0..xa = c; b = d; !!x. \ c \ x; x < d \ \ f x = g x \ \ setsum f {a..i \ Suc n. f i) = (\i \ n. f i) + f(Suc n)" by (simp add:atMost_Suc add_ac) lemma setsum_lessThan_Suc[simp]: "(\i < Suc n. f i) = (\i < n. f i) + f n" by (simp add:lessThan_Suc add_ac) lemma setsum_cl_ivl_Suc[simp]: "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))" by (auto simp:add_ac atLeastAtMostSuc_conv) lemma setsum_op_ivl_Suc[simp]: "setsum f {m.. (\i=n..m+1. f i) = (\i=n..m. f i) + f(m + 1)" by (auto simp:add_ac atLeastAtMostSuc_conv) *) lemma setsum_head: fixes n :: nat assumes mn: "m <= n" shows "(\x\{m..n}. P x) = P m + (\x\{m<..n}. P x)" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = (\x\{m} \ {m<..n}. P x)" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma setsum_head_Suc: "m \ n \ setsum f {m..n} = f m + setsum f {Suc m..n}" by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost) lemma setsum_head_upt_Suc: "m < n \ setsum f {m.. n + 1" shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using `m \ n+1` by auto thus ?thesis by (auto simp: ivl_disj_int setsum_Un_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma setsum_add_nat_ivl: "\ m \ n; n \ p \ \ setsum f {m.. 'a::ab_group_add" shows "\ m \ n; n \ p \ \ setsum f {m.. ('a::ab_group_add)" shows "setsum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m <= n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma setsum_restrict_set': "finite A \ setsum f {x \ A. x \ B} = (\x\A. if x \ B then f x else 0)" by (simp add: setsum_restrict_set [symmetric] Int_def) lemma setsum_restrict_set'': "finite A \ setsum f {x \ A. P x} = (\x\A. if P x then f x else 0)" by (simp add: setsum_restrict_set' [of A f "{x. P x}", simplified mem_Collect_eq]) lemma setsum_setsum_restrict: "finite S \ finite T \ setsum (\x. setsum (\y. f x y) {y. y \ T \ R x y}) S = setsum (\y. setsum (\x. f x y) {x. x \ S \ R x y}) T" by (simp add: setsum_restrict_set'') (rule setsum_commute) lemma setsum_image_gen: assumes fS: "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" by simp also have "\ = setsum (\y. setsum g {x. x \ S \ f x = y}) (f ` S)" by (rule setsum_setsum_restrict[OF fS finite_imageI[OF fS]]) finally show ?thesis . qed lemma setsum_le_included: fixes f :: "'a \ 'b::ordered_comm_monoid_add" 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" 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)" using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg) also have "... \ setsum g t" using assms by (auto simp: setsum_image_gen[symmetric]) finally show ?thesis . qed 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") proof- have "?l = setsum (\i. setsum (\x.1) {j\t. R i j}) s" by auto also have "\ = ?r" unfolding setsum_setsum_restrict[OF assms(1-2)] using assms(3) by auto finally show ?thesis . qed lemma setsum_multicount: 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) also have "\ = ?r" by(simp add: mult_commute) finally show ?thesis by auto qed subsection{* Shifting bounds *} lemma setsum_shift_bounds_nat_ivl: "setsum f {m+k.. setsum f {Suc 0..k} = setsum f {0..k}" by(simp add:setsum_head_Suc) lemma setsum_shift_lb_Suc0_0_upt: "f(0::nat) = 0 \ setsum f {Suc 0.. 1" shows "(\i=0.. 0" by simp_all moreover have "(\i=0.. 0` have "(1 + y) ^ n = (y * inverse y) * (1 + y) ^ n" by simp ultimately show ?case by (simp add: field_simps divide_inverse) qed ultimately show ?thesis by simp qed subsection {* The formula for arithmetic sums *} lemma gauss_sum: "(2::'a::comm_semiring_1)*(\i\{1..n}. of_nat i) = of_nat n*((of_nat n)+1)" proof (induct n) case 0 show ?case by simp next case (Suc n) then show ?case by (simp add: algebra_simps add: one_add_one [symmetric] del: one_add_one) (* FIXME: make numeral cancellation simprocs work for semirings *) qed theorem arith_series_general: "(2::'a::comm_semiring_1) * (\i\{.. 1" let ?I = "\i. of_nat i" and ?n = "of_nat n" have "(\i\{..i\{..i\{.. = ?n*a + (\i\{.. = (?n*a + d*(\i\{1.. = 2*?n*a + d*2*(\i\{1..i\{1..n - 1}. ?I i) = (2*?n*a + d*?I (n - 1)*?I n)" by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def) (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) finally show ?thesis unfolding mult_2 by (simp add: algebra_simps) next assume "\(n > 1)" hence "n = 1 \ n = 0" by auto thus ?thesis by (auto simp: mult_2) qed lemma arith_series_nat: "(2::nat) * (\i\{..i\{..i\{..nat" shows "\x. Q x \ P x \ (\xxxxx 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (xsymbols) "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) syntax (HTML output) "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10) "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ < _\<^raw:}$> _)" [0,0,10] 10) "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^raw:$\prod_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST setprod (%x. t) {a..b}" "\x=a..i\n. t" == "CONST setprod (\i. t) {..n}" "\ii. t) {..= 0 \ nat_set {x..y}" by (simp add: nat_set_def) declare transfer_morphism_nat_int[transfer add return: transfer_nat_int_set_functions transfer_nat_int_set_function_closures ] lemma transfer_int_nat_set_functions: "is_nat m \ is_nat n \ {m..n} = int ` {nat m..nat n}" by (simp only: is_nat_def transfer_nat_int_set_functions transfer_nat_int_set_function_closures transfer_nat_int_set_return_embed nat_0_le cong: transfer_nat_int_set_cong) lemma transfer_int_nat_set_function_closures: "is_nat x \ nat_set {x..y}" by (simp only: transfer_nat_int_set_function_closures is_nat_def) declare transfer_morphism_int_nat[transfer add return: transfer_int_nat_set_functions transfer_int_nat_set_function_closures ] end