huffman@47317: (* Title: HOL/Set_Interval.thy wenzelm@32960: Author: Tobias Nipkow wenzelm@32960: Author: Clemens Ballarin wenzelm@32960: Author: Jeremy Avigad nipkow@8924: ballarin@13735: lessThan, greaterThan, atLeast, atMost and two-sided intervals nipkow@51334: nipkow@51334: Modern convention: Ixy stands for an interval where x and y nipkow@51334: describe the lower and upper bound and x,y : {c,o,i} nipkow@51334: where c = closed, o = open, i = infinite. nipkow@51334: Examples: Ico = {_ ..< _} and Ici = {_ ..} nipkow@8924: *) nipkow@8924: wenzelm@60758: section \Set intervals\ wenzelm@14577: huffman@47317: theory Set_Interval blanchet@55088: imports Lattices_Big Nat_Transfer nipkow@15131: begin nipkow@8924: nipkow@24691: context ord nipkow@24691: begin haftmann@44008: nipkow@24691: definition wenzelm@32960: lessThan :: "'a => 'a set" ("(1{..<_})") where haftmann@25062: "{.. 'a set" ("(1{.._})") where haftmann@25062: "{..u} == {x. x \ u}" nipkow@24691: nipkow@24691: definition wenzelm@32960: greaterThan :: "'a => 'a set" ("(1{_<..})") where haftmann@25062: "{l<..} == {x. l 'a set" ("(1{_..})") where haftmann@25062: "{l..} == {x. l\x}" nipkow@24691: nipkow@24691: definition haftmann@25062: greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where haftmann@25062: "{l<.. 'a => 'a set" ("(1{_..<_})") where haftmann@25062: "{l.. 'a => 'a set" ("(1{_<.._})") where haftmann@25062: "{l<..u} == {l<..} Int {..u}" nipkow@24691: nipkow@24691: definition haftmann@25062: atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where haftmann@25062: "{l..u} == {l..} Int {..u}" nipkow@24691: nipkow@24691: end nipkow@8924: ballarin@13735: wenzelm@60758: text\A note of warning when using @{term"{.. nipkow@15048: kleing@14418: syntax huffman@36364: "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) huffman@36364: "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) huffman@36364: "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) huffman@36364: "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) kleing@14418: nipkow@30372: syntax (xsymbols) wenzelm@60586: "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) wenzelm@60586: "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) wenzelm@60586: "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) wenzelm@60586: "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) kleing@14418: nipkow@30372: syntax (latex output) huffman@36364: "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) huffman@36364: "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) huffman@36364: "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ \ _)/ _)" [0, 0, 10] 10) huffman@36364: "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(00_ < _)/ _)" [0, 0, 10] 10) kleing@14418: kleing@14418: translations kleing@14418: "UN i<=n. A" == "UN i:{..n}. A" nipkow@15045: "UN iVarious equivalences\ ballarin@13735: haftmann@25062: lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i { b <..} = { max a b <..}" hoelzl@50999: by auto hoelzl@50999: hoelzl@50999: lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" hoelzl@50999: by auto paulson@13850: wenzelm@60758: subsection \Logical Equivalences for Set Inclusion and Equality\ paulson@13850: paulson@13850: lemma atLeast_subset_iff [iff]: paulson@15418: "(atLeast x \ atLeast y) = (y \ (x::'a::order))" paulson@15418: by (blast intro: order_trans) paulson@13850: paulson@13850: lemma atLeast_eq_iff [iff]: paulson@15418: "(atLeast x = atLeast y) = (x = (y::'a::linorder))" paulson@13850: by (blast intro: order_antisym order_trans) paulson@13850: paulson@13850: lemma greaterThan_subset_iff [iff]: paulson@15418: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" paulson@15418: apply (auto simp add: greaterThan_def) paulson@15418: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma greaterThan_eq_iff [iff]: paulson@15418: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" paulson@15418: apply (rule iffI) paulson@15418: apply (erule equalityE) haftmann@29709: apply simp_all paulson@13850: done paulson@13850: paulson@15418: lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::order))" paulson@13850: by (blast intro: order_trans) paulson@13850: paulson@15418: lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))" paulson@13850: by (blast intro: order_antisym order_trans) paulson@13850: paulson@13850: lemma lessThan_subset_iff [iff]: paulson@15418: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" paulson@15418: apply (auto simp add: lessThan_def) paulson@15418: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma lessThan_eq_iff [iff]: paulson@15418: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" paulson@15418: apply (rule iffI) paulson@15418: apply (erule equalityE) haftmann@29709: apply simp_all ballarin@13735: done ballarin@13735: hoelzl@40703: lemma lessThan_strict_subset_iff: hoelzl@40703: fixes m n :: "'a::linorder" hoelzl@40703: shows "{.. m < n" hoelzl@40703: by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) ballarin@13735: hoelzl@57448: lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" hoelzl@57448: by auto hoelzl@57448: hoelzl@57448: lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" hoelzl@57448: by auto hoelzl@57448: wenzelm@60758: subsection \Two-sided intervals\ ballarin@13735: nipkow@24691: context ord nipkow@24691: begin nipkow@24691: blanchet@54147: lemma greaterThanLessThan_iff [simp]: haftmann@25062: "(i : {l<..The above four lemmas could be declared as iffs. Unfortunately this haftmann@52729: breaks many proofs. Since it only helps blast, it is better to leave them wenzelm@60758: alone.\ nipkow@32436: hoelzl@50999: lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" hoelzl@50999: by auto hoelzl@50999: nipkow@24691: end ballarin@13735: wenzelm@60758: subsubsection\Emptyness, singletons, subset\ nipkow@15554: nipkow@24691: context order nipkow@24691: begin nipkow@15554: nipkow@32400: lemma atLeastatMost_empty[simp]: nipkow@32400: "b < a \ {a..b} = {}" nipkow@32400: by(auto simp: atLeastAtMost_def atLeast_def atMost_def) nipkow@32400: nipkow@32400: lemma atLeastatMost_empty_iff[simp]: nipkow@32400: "{a..b} = {} \ (~ a <= b)" nipkow@32400: by auto (blast intro: order_trans) nipkow@32400: nipkow@32400: lemma atLeastatMost_empty_iff2[simp]: nipkow@32400: "{} = {a..b} \ (~ a <= b)" nipkow@32400: by auto (blast intro: order_trans) nipkow@32400: nipkow@32400: lemma atLeastLessThan_empty[simp]: nipkow@32400: "b <= a \ {a.. (~ a < b)" nipkow@32400: by auto (blast intro: le_less_trans) nipkow@32400: nipkow@32400: lemma atLeastLessThan_empty_iff2[simp]: nipkow@32400: "{} = {a.. (~ a < b)" nipkow@32400: by auto (blast intro: le_less_trans) nipkow@15554: nipkow@32400: lemma greaterThanAtMost_empty[simp]: "l \ k ==> {k<..l} = {}" nipkow@17719: by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) nipkow@17719: nipkow@32400: lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ ~ k < l" nipkow@32400: by auto (blast intro: less_le_trans) nipkow@32400: nipkow@32400: lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ ~ k < l" nipkow@32400: by auto (blast intro: less_le_trans) nipkow@32400: haftmann@29709: lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp hoelzl@36846: nipkow@32400: lemma atLeastatMost_subset_iff[simp]: nipkow@32400: "{a..b} <= {c..d} \ (~ a <= b) | c <= a & b <= d" nipkow@32400: unfolding atLeastAtMost_def atLeast_def atMost_def nipkow@32400: by (blast intro: order_trans) nipkow@32400: nipkow@32400: lemma atLeastatMost_psubset_iff: nipkow@32400: "{a..b} < {c..d} \ nipkow@32400: ((~ a <= b) | c <= a & b <= d & (c < a | b < d)) & c <= d" nipkow@39302: by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) nipkow@32400: nipkow@51334: lemma Icc_eq_Icc[simp]: nipkow@51334: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" nipkow@51334: by(simp add: order_class.eq_iff)(auto intro: order_trans) nipkow@51334: hoelzl@36846: lemma atLeastAtMost_singleton_iff[simp]: hoelzl@36846: "{a .. b} = {c} \ a = b \ b = c" hoelzl@36846: proof hoelzl@36846: assume "{a..b} = {c}" wenzelm@53374: hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp wenzelm@60758: with \{a..b} = {c}\ have "c \ a \ b \ c" by auto wenzelm@53374: with * show "a = b \ b = c" by auto hoelzl@36846: qed simp hoelzl@36846: nipkow@51334: lemma Icc_subset_Ici_iff[simp]: nipkow@51334: "{l..h} \ {l'..} = (~ l\h \ l\l')" nipkow@51334: by(auto simp: subset_eq intro: order_trans) nipkow@51334: nipkow@51334: lemma Icc_subset_Iic_iff[simp]: nipkow@51334: "{l..h} \ {..h'} = (~ l\h \ h\h')" nipkow@51334: by(auto simp: subset_eq intro: order_trans) nipkow@51334: nipkow@51334: lemma not_Ici_eq_empty[simp]: "{l..} \ {}" nipkow@51334: by(auto simp: set_eq_iff) nipkow@51334: nipkow@51334: lemma not_Iic_eq_empty[simp]: "{..h} \ {}" nipkow@51334: by(auto simp: set_eq_iff) nipkow@51334: nipkow@51334: lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] nipkow@51334: lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] nipkow@51334: nipkow@24691: end paulson@14485: nipkow@51334: context no_top nipkow@51334: begin nipkow@51334: nipkow@51334: (* also holds for no_bot but no_top should suffice *) nipkow@51334: lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" nipkow@51334: using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) nipkow@51334: nipkow@51334: lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" nipkow@51334: using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) nipkow@51334: nipkow@51334: lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" nipkow@51334: using gt_ex[of h'] nipkow@51334: by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) nipkow@51334: nipkow@51334: lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" nipkow@51334: using gt_ex[of h'] nipkow@51334: by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) nipkow@51334: nipkow@51334: end nipkow@51334: nipkow@51334: context no_bot nipkow@51334: begin nipkow@51334: nipkow@51334: lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" nipkow@51334: using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) nipkow@51334: nipkow@51334: lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" nipkow@51334: using lt_ex[of l'] nipkow@51334: by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) nipkow@51334: nipkow@51334: lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" nipkow@51334: using lt_ex[of l'] nipkow@51334: by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) nipkow@51334: nipkow@51334: end nipkow@51334: nipkow@51334: nipkow@51334: context no_top nipkow@51334: begin nipkow@51334: nipkow@51334: (* also holds for no_bot but no_top should suffice *) nipkow@51334: lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" nipkow@51334: using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) nipkow@51334: nipkow@51334: lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] nipkow@51334: nipkow@51334: lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" nipkow@51334: using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) nipkow@51334: nipkow@51334: lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] nipkow@51334: nipkow@51334: lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" nipkow@51334: unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast nipkow@51334: nipkow@51334: lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] nipkow@51334: nipkow@51334: (* also holds for no_bot but no_top should suffice *) nipkow@51334: lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" nipkow@51334: using not_Ici_le_Iic[of l' h] by blast nipkow@51334: nipkow@51334: lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] nipkow@51334: nipkow@51334: end nipkow@51334: nipkow@51334: context no_bot nipkow@51334: begin nipkow@51334: nipkow@51334: lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" nipkow@51334: using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) nipkow@51334: nipkow@51334: lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] nipkow@51334: nipkow@51334: lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" nipkow@51334: unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast nipkow@51334: nipkow@51334: lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] nipkow@51334: nipkow@51334: end nipkow@51334: nipkow@51334: hoelzl@53216: context dense_linorder hoelzl@42891: begin hoelzl@42891: hoelzl@42891: lemma greaterThanLessThan_empty_iff[simp]: hoelzl@42891: "{ a <..< b } = {} \ b \ a" hoelzl@42891: using dense[of a b] by (cases "a < b") auto hoelzl@42891: hoelzl@42891: lemma greaterThanLessThan_empty_iff2[simp]: hoelzl@42891: "{} = { a <..< b } \ b \ a" hoelzl@42891: using dense[of a b] by (cases "a < b") auto hoelzl@42891: hoelzl@42901: lemma atLeastLessThan_subseteq_atLeastAtMost_iff: hoelzl@42901: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" hoelzl@42901: using dense[of "max a d" "b"] hoelzl@42901: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@42901: hoelzl@42901: lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: hoelzl@42901: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" hoelzl@42901: using dense[of "a" "min c b"] hoelzl@42901: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@42901: hoelzl@42901: lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: hoelzl@42901: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" hoelzl@42901: using dense[of "a" "min c b"] dense[of "max a d" "b"] hoelzl@42901: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@42901: hoelzl@43657: lemma atLeastAtMost_subseteq_atLeastLessThan_iff: hoelzl@43657: "{a .. b} \ { c ..< d } \ (a \ b \ c \ a \ b < d)" hoelzl@43657: using dense[of "max a d" "b"] hoelzl@43657: by (force simp: subset_eq Ball_def not_less[symmetric]) eberlm@61524: eberlm@61524: lemma greaterThanLessThan_subseteq_greaterThanLessThan: eberlm@61524: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" eberlm@61524: using dense[of "a" "min c b"] dense[of "max a d" "b"] eberlm@61524: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@43657: hoelzl@43657: lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: hoelzl@43657: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" hoelzl@43657: using dense[of "a" "min c b"] hoelzl@43657: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@43657: hoelzl@43657: lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: hoelzl@43657: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" hoelzl@43657: using dense[of "a" "min c b"] dense[of "max a d" "b"] hoelzl@43657: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@43657: hoelzl@56328: lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: hoelzl@56328: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" hoelzl@56328: using dense[of "a" "min c b"] dense[of "max a d" "b"] hoelzl@56328: by (force simp: subset_eq Ball_def not_less[symmetric]) hoelzl@56328: hoelzl@42891: end hoelzl@42891: hoelzl@51329: context no_top hoelzl@51329: begin hoelzl@51329: nipkow@51334: lemma greaterThan_non_empty[simp]: "{x <..} \ {}" hoelzl@51329: using gt_ex[of x] by auto hoelzl@51329: hoelzl@51329: end hoelzl@51329: hoelzl@51329: context no_bot hoelzl@51329: begin hoelzl@51329: nipkow@51334: lemma lessThan_non_empty[simp]: "{..< x} \ {}" hoelzl@51329: using lt_ex[of x] by auto hoelzl@51329: hoelzl@51329: end hoelzl@51329: nipkow@32408: lemma (in linorder) atLeastLessThan_subset_iff: nipkow@32408: "{a.. b <= a | c<=a & b<=d" nipkow@32408: apply (auto simp:subset_eq Ball_def) nipkow@32408: apply(frule_tac x=a in spec) nipkow@32408: apply(erule_tac x=d in allE) nipkow@32408: apply (simp add: less_imp_le) nipkow@32408: done nipkow@32408: hoelzl@40703: lemma atLeastLessThan_inj: hoelzl@40703: fixes a b c d :: "'a::linorder" hoelzl@40703: assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" hoelzl@40703: shows "a = c" "b = d" hoelzl@40703: using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le linorder_antisym_conv2 subset_refl)+ hoelzl@40703: hoelzl@40703: lemma atLeastLessThan_eq_iff: hoelzl@40703: fixes a b c d :: "'a::linorder" hoelzl@40703: assumes "a < b" "c < d" hoelzl@40703: shows "{a ..< b} = {c ..< d} \ a = c \ b = d" hoelzl@40703: using atLeastLessThan_inj assms by auto hoelzl@40703: hoelzl@57447: lemma (in linorder) Ioc_inj: "{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d" hoelzl@57447: by (metis eq_iff greaterThanAtMost_empty_iff2 greaterThanAtMost_iff le_cases not_le) hoelzl@57447: hoelzl@57447: lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" hoelzl@57447: by auto hoelzl@57447: hoelzl@57447: lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" hoelzl@57447: by (auto simp: subset_eq Ball_def) (metis less_le not_less) hoelzl@57447: haftmann@52729: lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" nipkow@51334: by (auto simp: set_eq_iff intro: le_bot) hoelzl@51328: haftmann@52729: lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" nipkow@51334: by (auto simp: set_eq_iff intro: top_le) hoelzl@51328: nipkow@51334: lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: nipkow@51334: "{x..y} = UNIV \ (x = bot \ y = top)" nipkow@51334: by (auto simp: set_eq_iff intro: top_le le_bot) hoelzl@51328: hoelzl@56949: lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" hoelzl@56949: by (auto simp: set_eq_iff not_less le_bot) hoelzl@56949: hoelzl@56949: lemma Iio_eq_empty_iff_nat: "{..< n::nat} = {} \ n = 0" hoelzl@56949: by (simp add: Iio_eq_empty_iff bot_nat_def) hoelzl@56949: noschinl@58970: lemma mono_image_least: noschinl@58970: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" noschinl@58970: shows "f m = m'" noschinl@58970: proof - noschinl@58970: from f_img have "{m' ..< n'} \ {}" noschinl@58970: by (metis atLeastLessThan_empty_iff image_is_empty) noschinl@58970: with f_img have "m' \ f ` {m ..< n}" by auto noschinl@58970: then obtain k where "f k = m'" "m \ k" by auto noschinl@58970: moreover have "m' \ f m" using f_img by auto noschinl@58970: ultimately show "f m = m'" noschinl@58970: using f_mono by (auto elim: monoE[where x=m and y=k]) noschinl@58970: qed noschinl@58970: hoelzl@51328: wenzelm@60758: subsection \Infinite intervals\ hoelzl@56328: hoelzl@56328: context dense_linorder hoelzl@56328: begin hoelzl@56328: hoelzl@56328: lemma infinite_Ioo: hoelzl@56328: assumes "a < b" hoelzl@56328: shows "\ finite {a<.. {}" wenzelm@60758: using \a < b\ by auto hoelzl@56328: ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" hoelzl@56328: using Max_in[of "{a <..< b}"] by auto hoelzl@56328: then obtain x where "Max {a <..< b} < x" "x < b" hoelzl@56328: using dense[of "Max {a<.. {a <..< b}" wenzelm@60758: using \a < Max {a <..< b}\ by auto hoelzl@56328: then have "x \ Max {a <..< b}" hoelzl@56328: using fin by auto wenzelm@60758: with \Max {a <..< b} < x\ show False by auto hoelzl@56328: qed hoelzl@56328: hoelzl@56328: lemma infinite_Icc: "a < b \ \ finite {a .. b}" hoelzl@56328: using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] hoelzl@56328: by (auto dest: finite_subset) hoelzl@56328: hoelzl@56328: lemma infinite_Ico: "a < b \ \ finite {a ..< b}" hoelzl@56328: using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] hoelzl@56328: by (auto dest: finite_subset) hoelzl@56328: hoelzl@56328: lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" hoelzl@56328: using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] hoelzl@56328: by (auto dest: finite_subset) hoelzl@56328: hoelzl@56328: end hoelzl@56328: hoelzl@56328: lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" hoelzl@56328: proof hoelzl@56328: assume "finite {..< a}" hoelzl@56328: then have *: "\x. x < a \ Min {..< a} \ x" hoelzl@56328: by auto hoelzl@56328: obtain x where "x < a" hoelzl@56328: using lt_ex by auto hoelzl@56328: hoelzl@56328: obtain y where "y < Min {..< a}" hoelzl@56328: using lt_ex by auto hoelzl@56328: also have "Min {..< a} \ x" wenzelm@60758: using \x < a\ by fact wenzelm@60758: also note \x < a\ hoelzl@56328: finally have "Min {..< a} \ y" hoelzl@56328: by fact wenzelm@60758: with \y < Min {..< a}\ show False by auto hoelzl@56328: qed hoelzl@56328: hoelzl@56328: lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" hoelzl@56328: using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] hoelzl@56328: by (auto simp: subset_eq less_imp_le) hoelzl@56328: hoelzl@56328: lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" hoelzl@56328: proof hoelzl@56328: assume "finite {a <..}" hoelzl@56328: then have *: "\x. a < x \ x \ Max {a <..}" hoelzl@56328: by auto hoelzl@56328: hoelzl@56328: obtain y where "Max {a <..} < y" hoelzl@56328: using gt_ex by auto hoelzl@56328: hoelzl@56328: obtain x where "a < x" hoelzl@56328: using gt_ex by auto hoelzl@56328: also then have "x \ Max {a <..}" hoelzl@56328: by fact wenzelm@60758: also note \Max {a <..} < y\ hoelzl@56328: finally have "y \ Max { a <..}" hoelzl@56328: by fact wenzelm@60758: with \Max {a <..} < y\ show False by auto hoelzl@56328: qed hoelzl@56328: hoelzl@56328: lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" hoelzl@56328: using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] hoelzl@56328: by (auto simp: subset_eq less_imp_le) hoelzl@56328: wenzelm@60758: subsubsection \Intersection\ nipkow@32456: nipkow@32456: context linorder nipkow@32456: begin nipkow@32456: nipkow@32456: lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" nipkow@32456: by auto nipkow@32456: nipkow@32456: lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" hoelzl@50417: by (auto simp: min_def) hoelzl@50417: hoelzl@57447: lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" hoelzl@57447: using assms by auto hoelzl@57447: nipkow@32456: end nipkow@32456: hoelzl@51329: context complete_lattice hoelzl@51329: begin hoelzl@51329: hoelzl@51329: lemma hoelzl@51329: shows Sup_atLeast[simp]: "Sup {x ..} = top" hoelzl@51329: and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" hoelzl@51329: and Sup_atMost[simp]: "Sup {.. y} = y" hoelzl@51329: and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" hoelzl@51329: and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" hoelzl@51329: by (auto intro!: Sup_eqI) hoelzl@51329: hoelzl@51329: lemma hoelzl@51329: shows Inf_atMost[simp]: "Inf {.. x} = bot" hoelzl@51329: and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" hoelzl@51329: and Inf_atLeast[simp]: "Inf {x ..} = x" hoelzl@51329: and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" hoelzl@51329: and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" hoelzl@51329: by (auto intro!: Inf_eqI) hoelzl@51329: hoelzl@51329: end hoelzl@51329: hoelzl@51329: lemma hoelzl@53216: fixes x y :: "'a :: {complete_lattice, dense_linorder}" hoelzl@51329: shows Sup_lessThan[simp]: "Sup {..< y} = y" hoelzl@51329: and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" hoelzl@51329: and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" hoelzl@51329: and Inf_greaterThan[simp]: "Inf {x <..} = x" hoelzl@51329: and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" hoelzl@51329: and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" hoelzl@51329: by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) nipkow@32456: wenzelm@60758: subsection \Intervals of natural numbers\ paulson@14485: wenzelm@60758: subsubsection \The Constant @{term lessThan}\ paulson@15047: paulson@14485: lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" paulson@14485: by (simp add: lessThan_def) paulson@14485: paulson@14485: lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" paulson@14485: by (simp add: lessThan_def less_Suc_eq, blast) paulson@14485: wenzelm@60758: text \The following proof is convenient in induction proofs where hoelzl@39072: new elements get indices at the beginning. So it is used to transform wenzelm@60758: @{term "{.. hoelzl@39072: hoelzl@59000: lemma zero_notin_Suc_image: "0 \ Suc ` A" hoelzl@59000: by auto hoelzl@59000: hoelzl@39072: lemma lessThan_Suc_eq_insert_0: "{..The Constant @{term greaterThan}\ paulson@15047: paulson@14485: lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc" paulson@14485: apply (simp add: greaterThan_def) paulson@14485: apply (blast dest: gr0_conv_Suc [THEN iffD1]) paulson@14485: done paulson@14485: paulson@14485: lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" paulson@14485: apply (simp add: greaterThan_def) paulson@14485: apply (auto elim: linorder_neqE) paulson@14485: done paulson@14485: paulson@14485: lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}" paulson@14485: by blast paulson@14485: wenzelm@60758: subsubsection \The Constant @{term atLeast}\ paulson@15047: paulson@14485: lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" paulson@14485: by (unfold atLeast_def UNIV_def, simp) paulson@14485: paulson@14485: lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" paulson@14485: apply (simp add: atLeast_def) paulson@14485: apply (simp add: Suc_le_eq) paulson@14485: apply (simp add: order_le_less, blast) paulson@14485: done paulson@14485: paulson@14485: lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" paulson@14485: by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) paulson@14485: paulson@14485: lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV" paulson@14485: by blast paulson@14485: wenzelm@60758: subsubsection \The Constant @{term atMost}\ paulson@15047: paulson@14485: lemma atMost_0 [simp]: "atMost (0::nat) = {0}" paulson@14485: by (simp add: atMost_def) paulson@14485: paulson@14485: lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" paulson@14485: apply (simp add: atMost_def) paulson@14485: apply (simp add: less_Suc_eq order_le_less, blast) paulson@14485: done paulson@14485: paulson@14485: lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV" paulson@14485: by blast paulson@14485: wenzelm@60758: subsubsection \The Constant @{term atLeastLessThan}\ paulson@15047: wenzelm@60758: text\The orientation of the following 2 rules is tricky. The lhs is nipkow@24449: defined in terms of the rhs. Hence the chosen orientation makes sense nipkow@24449: in this theory --- the reverse orientation complicates proofs (eg nipkow@24449: nontermination). But outside, when the definition of the lhs is rarely nipkow@24449: used, the opposite orientation seems preferable because it reduces a wenzelm@60758: specific concept to a more general one.\ nipkow@28068: paulson@15047: lemma atLeast0LessThan: "{0::nat..Intervals of nats with @{term Suc}\ paulson@15047: wenzelm@60758: text\Not a simprule because the RHS is too messy.\ paulson@15047: lemma atLeastLessThanSuc: paulson@15047: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" nipkow@15554: by (auto simp add: atLeastAtMost_def) nipkow@15554: noschinl@45932: lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" noschinl@45932: by auto noschinl@45932: wenzelm@60758: text \The analogous result is useful on @{typ int}:\ kleing@43157: (* here, because we don't have an own int section *) kleing@43157: lemma atLeastAtMostPlus1_int_conv: kleing@43157: "m <= 1+n \ {m..1+n} = insert (1+n) {m..n::int}" kleing@43157: by (auto intro: set_eqI) kleing@43157: paulson@33044: lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lp15@57113: wenzelm@61799: lemma lessThan_nat_numeral: \\Evaluation for specific numerals\ lp15@57113: "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" lp15@57113: by (simp add: numeral_eq_Suc lessThan_Suc) lp15@57113: wenzelm@61799: lemma atMost_nat_numeral: \\Evaluation for specific numerals\ lp15@57113: "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" lp15@57113: by (simp add: numeral_eq_Suc atMost_Suc) lp15@57113: wenzelm@61799: lemma atLeastLessThan_nat_numeral: \\Evaluation for specific numerals\ lp15@57113: "atLeastLessThan m (numeral k :: nat) = lp15@57113: (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) lp15@57113: else {})" lp15@57113: by (simp add: numeral_eq_Suc atLeastLessThanSuc) lp15@57113: wenzelm@60758: subsubsection \Image\ nipkow@16733: lp15@60809: lemma image_add_atLeastAtMost [simp]: lp15@60615: fixes k ::"'a::linordered_semidom" lp15@60615: shows "(\n. n+k) ` {i..j} = {i+k..j+k}" (is "?A = ?B") nipkow@16733: proof nipkow@16733: show "?A \ ?B" by auto nipkow@16733: next nipkow@16733: show "?B \ ?A" nipkow@16733: proof nipkow@16733: fix n assume a: "n : ?B" lp15@60615: hence "n - k : {i..j}" lp15@60615: by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) lp15@60615: moreover have "n = (n - k) + k" using a lp15@60615: proof - lp15@60615: have "k + i \ n" lp15@60615: by (metis a add.commute atLeastAtMost_iff) lp15@60615: hence "k + (n - k) = n" lp15@60615: by (metis (no_types) ab_semigroup_add_class.add_ac(1) add_diff_cancel_left' le_add_diff_inverse) lp15@60615: thus ?thesis lp15@60615: by (simp add: add.commute) lp15@60615: qed nipkow@16733: ultimately show "n : ?A" by blast nipkow@16733: qed nipkow@16733: qed nipkow@16733: lp15@60809: lemma image_diff_atLeastAtMost [simp]: lp15@60809: fixes d::"'a::linordered_idom" shows "(op - d ` {a..b}) = {d-b..d-a}" lp15@60809: apply auto lp15@60809: apply (rule_tac x="d-x" in rev_image_eqI, auto) lp15@60809: done lp15@60809: lp15@60809: lemma image_mult_atLeastAtMost [simp]: lp15@60809: fixes d::"'a::linordered_field" lp15@60809: assumes "d>0" shows "(op * d ` {a..b}) = {d*a..d*b}" lp15@60809: using assms lp15@60809: by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lp15@60809: lp15@60809: lemma image_affinity_atLeastAtMost: lp15@60809: fixes c :: "'a::linordered_field" lp15@60809: shows "((\x. m*x + c) ` {a..b}) = (if {a..b}={} then {} lp15@60809: else if 0 \ m then {m*a + c .. m *b + c} lp15@60809: else {m*b + c .. m*a + c})" lp15@60809: apply (case_tac "m=0", auto simp: mult_le_cancel_left) lp15@60809: apply (rule_tac x="inverse m*(x-c)" in rev_image_eqI, auto simp: field_simps) lp15@60809: apply (rule_tac x="inverse m*(x-c)" in rev_image_eqI, auto simp: field_simps) lp15@60809: done lp15@60809: lp15@60809: lemma image_affinity_atLeastAtMost_diff: lp15@60809: fixes c :: "'a::linordered_field" lp15@60809: shows "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} lp15@60809: else if 0 \ m then {m*a - c .. m*b - c} lp15@60809: else {m*b - c .. m*a - c})" lp15@60809: using image_affinity_atLeastAtMost [of m "-c" a b] lp15@60809: by simp lp15@60809: paulson@61204: lemma image_affinity_atLeastAtMost_div: paulson@61204: fixes c :: "'a::linordered_field" paulson@61204: shows "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} paulson@61204: else if 0 \ m then {a/m + c .. b/m + c} paulson@61204: else {b/m + c .. a/m + c})" paulson@61204: using image_affinity_atLeastAtMost [of "inverse m" c a b] paulson@61204: by (simp add: field_class.field_divide_inverse algebra_simps) paulson@61204: lp15@60809: lemma image_affinity_atLeastAtMost_div_diff: lp15@60809: fixes c :: "'a::linordered_field" lp15@60809: shows "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} lp15@60809: else if 0 \ m then {a/m - c .. b/m - c} lp15@60809: else {b/m - c .. a/m - c})" lp15@60809: using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] lp15@60809: by (simp add: field_class.field_divide_inverse algebra_simps) lp15@60809: nipkow@16733: lemma image_add_atLeastLessThan: nipkow@16733: "(%n::nat. n+k) ` {i.. ?B" by auto nipkow@16733: next nipkow@16733: show "?B \ ?A" nipkow@16733: proof nipkow@16733: fix n assume a: "n : ?B" webertj@20217: hence "n - k : {i..i. i - c) ` {x ..< y} = hoelzl@37664: (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" hoelzl@37664: (is "_ = ?right") hoelzl@37664: proof safe hoelzl@37664: fix a assume a: "a \ ?right" hoelzl@37664: show "a \ (\i. i - c) ` {x ..< y}" hoelzl@37664: proof cases hoelzl@37664: assume "c < y" with a show ?thesis hoelzl@37664: by (auto intro!: image_eqI[of _ _ "a + c"]) hoelzl@37664: next hoelzl@37664: assume "\ c < y" with a show ?thesis hoelzl@37664: by (auto intro!: image_eqI[of _ _ x] split: split_if_asm) hoelzl@37664: qed hoelzl@37664: qed auto hoelzl@37664: Andreas@51152: lemma image_int_atLeastLessThan: "int ` {a.. uminus ` {x<..}" hoelzl@35580: by (rule imageI) (simp add: *) hoelzl@35580: thus "y \ uminus ` {x<..}" by simp hoelzl@35580: next hoelzl@35580: fix y assume "y \ -x" hoelzl@35580: have "- (-y) \ uminus ` {x..}" wenzelm@60758: by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) hoelzl@35580: thus "y \ uminus ` {x..}" by simp hoelzl@35580: qed simp_all hoelzl@35580: hoelzl@35580: lemma hoelzl@35580: fixes x :: 'a hoelzl@35580: shows image_uminus_lessThan[simp]: "uminus ` {..Finiteness\ paulson@14485: nipkow@15045: lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ paulson@14485: lemma bounded_nat_set_is_finite: nipkow@24853: "(ALL i:N. i < (n::nat)) ==> finite N" nipkow@28068: apply (rule finite_subset) nipkow@28068: apply (rule_tac [2] finite_lessThan, auto) nipkow@28068: done nipkow@28068: wenzelm@60758: text \A set of natural numbers is finite iff it is bounded.\ nipkow@31044: lemma finite_nat_set_iff_bounded: nipkow@31044: "finite(N::nat set) = (EX m. ALL n:N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast nipkow@31044: next wenzelm@60758: assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) nipkow@31044: qed nipkow@31044: nipkow@31044: lemma finite_nat_set_iff_bounded_le: nipkow@31044: "finite(N::nat set) = (EX m. ALL n:N. n<=m)" nipkow@31044: apply(simp add:finite_nat_set_iff_bounded) nipkow@31044: apply(blast dest:less_imp_le_nat le_imp_less_Suc) nipkow@31044: done nipkow@31044: nipkow@28068: lemma finite_less_ub: nipkow@28068: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" nipkow@28068: by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) paulson@14485: hoelzl@56328: wenzelm@60758: text\Any subset of an interval of natural numbers the size of the wenzelm@60758: subset is exactly that interval.\ nipkow@24853: nipkow@24853: lemma subset_card_intvl_is_intvl: blanchet@55085: assumes "A \ {k.. A" by auto blanchet@55085: with insert have "A <= {k..Proving Inclusions and Equalities between Unions\ paulson@32596: nipkow@36755: lemma UN_le_eq_Un0: nipkow@36755: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") nipkow@36755: proof nipkow@36755: show "?A <= ?B" nipkow@36755: proof nipkow@36755: fix x assume "x : ?A" nipkow@36755: then obtain i where i: "i\n" "x : M i" by auto nipkow@36755: show "x : ?B" nipkow@36755: proof(cases i) nipkow@36755: case 0 with i show ?thesis by simp nipkow@36755: next nipkow@36755: case (Suc j) with i show ?thesis by auto nipkow@36755: qed nipkow@36755: qed nipkow@36755: next nipkow@36755: show "?B <= ?A" by auto nipkow@36755: qed nipkow@36755: nipkow@36755: lemma UN_le_add_shift: nipkow@36755: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") nipkow@36755: proof nipkow@44890: show "?A <= ?B" by fastforce nipkow@36755: next nipkow@36755: show "?B <= ?A" nipkow@36755: proof nipkow@36755: fix x assume "x : ?B" nipkow@36755: then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto nipkow@36755: hence "i-k\n & x : M((i-k)+k)" by auto nipkow@36755: thus "x : ?A" by blast nipkow@36755: qed nipkow@36755: qed nipkow@36755: paulson@32596: lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" paulson@32596: by (auto simp add: atLeast0LessThan) paulson@32596: paulson@32596: lemma UN_finite_subset: "(!!n::nat. (\i\{0.. C) \ (\n. A n) \ C" paulson@32596: by (subst UN_UN_finite_eq [symmetric]) blast paulson@32596: paulson@33044: lemma UN_finite2_subset: paulson@33044: "(!!n::nat. (\i\{0.. (\i\{0.. (\n. A n) \ (\n. B n)" paulson@33044: apply (rule UN_finite_subset) paulson@33044: apply (subst UN_UN_finite_eq [symmetric, of B]) paulson@33044: apply blast paulson@33044: done paulson@32596: paulson@32596: lemma UN_finite2_eq: paulson@33044: "(!!n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" paulson@33044: apply (rule subset_antisym) paulson@33044: apply (rule UN_finite2_subset, blast) paulson@33044: apply (rule UN_finite2_subset [where k=k]) huffman@35216: apply (force simp add: atLeastLessThan_add_Un [of 0]) paulson@33044: done paulson@32596: paulson@32596: wenzelm@60758: subsubsection \Cardinality\ paulson@14485: nipkow@15045: 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" nipkow@31438: apply(drule ex_bij_betw_finite_nat) nipkow@31438: apply(drule ex_bij_betw_nat_finite) nipkow@31438: apply(auto intro!:bij_betw_trans) nipkow@31438: done nipkow@31438: nipkow@31438: lemma ex_bij_betw_nat_finite_1: nipkow@31438: "finite M \ \h. bij_betw h {1 .. card M} M" nipkow@31438: by (rule finite_same_card_bij) auto nipkow@31438: hoelzl@40703: lemma bij_betw_iff_card: hoelzl@40703: assumes FIN: "finite A" and FIN': "finite B" hoelzl@40703: shows BIJ: "(\f. bij_betw f A B) \ (card A = card B)" hoelzl@40703: using assms hoelzl@40703: proof(auto simp add: bij_betw_same_card) hoelzl@40703: assume *: "card A = card B" hoelzl@40703: obtain f where "bij_betw f A {0 ..< card A}" hoelzl@40703: using FIN ex_bij_betw_finite_nat by blast hoelzl@40703: moreover obtain g where "bij_betw g {0 ..< card B} B" hoelzl@40703: using FIN' ex_bij_betw_nat_finite by blast hoelzl@40703: ultimately have "bij_betw (g o f) A B" hoelzl@40703: using * by (auto simp add: bij_betw_trans) hoelzl@40703: thus "(\f. bij_betw f A B)" by blast hoelzl@40703: qed hoelzl@40703: hoelzl@40703: lemma inj_on_iff_card_le: hoelzl@40703: assumes FIN: "finite A" and FIN': "finite B" hoelzl@40703: shows "(\f. inj_on f A \ f ` A \ B) = (card A \ card B)" hoelzl@40703: proof (safe intro!: card_inj_on_le) hoelzl@40703: assume *: "card A \ card B" hoelzl@40703: obtain f where 1: "inj_on f A" and 2: "f ` A = {0 ..< card A}" hoelzl@40703: using FIN ex_bij_betw_finite_nat unfolding bij_betw_def by force hoelzl@40703: moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B" hoelzl@40703: using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force hoelzl@40703: ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force hoelzl@40703: hence "inj_on (g o f) A" using 1 comp_inj_on by blast hoelzl@40703: moreover hoelzl@40703: {have "{0 ..< card A} \ {0 ..< card B}" using * by force hoelzl@40703: with 2 have "f ` A \ {0 ..< card B}" by blast hoelzl@40703: hence "(g o f) ` A \ B" unfolding comp_def using 3 by force hoelzl@40703: } hoelzl@40703: ultimately show "(\f. inj_on f A \ f ` A \ B)" by blast hoelzl@40703: qed (insert assms, auto) nipkow@26105: wenzelm@60758: subsection \Intervals of integers\ paulson@14485: nipkow@15045: lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ paulson@14485: paulson@15418: lemma image_atLeastZeroLessThan_int: "0 \ u ==> nipkow@15045: {(0::int).. u") paulson@14485: apply (subst image_atLeastZeroLessThan_int, assumption) paulson@14485: apply (rule finite_imageI) paulson@14485: apply auto paulson@14485: done paulson@14485: nipkow@15045: lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ paulson@14485: nipkow@15045: lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") paulson@14485: apply (subst image_atLeastZeroLessThan_int, assumption) paulson@14485: apply (subst card_image) paulson@14485: apply (auto simp add: inj_on_def) paulson@14485: done paulson@14485: nipkow@15045: lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" bulwahn@27656: proof - bulwahn@27656: have "{k. P k \ k < i} \ {.. M" bulwahn@27656: shows "card {k \ M. k < Suc i} \ 0" bulwahn@27656: proof - bulwahn@27656: from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto bulwahn@27656: with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) bulwahn@27656: qed bulwahn@27656: bulwahn@27656: lemma card_less_Suc2: "0 \ M \ card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" haftmann@37388: apply (rule card_bij_eq [of Suc _ _ "\x. x - Suc 0"]) bulwahn@27656: apply auto bulwahn@27656: apply (rule inj_on_diff_nat) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac x) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac xa) bulwahn@27656: apply auto bulwahn@27656: apply (case_tac xa) bulwahn@27656: apply auto bulwahn@27656: done bulwahn@27656: bulwahn@27656: lemma card_less_Suc: bulwahn@27656: assumes zero_in_M: "0 \ M" bulwahn@27656: shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" bulwahn@27656: proof - bulwahn@27656: from assms have a: "0 \ {k \ M. k < Suc i}" by simp bulwahn@27656: hence c: "{k \ M. k < Suc i} = insert 0 ({k \ M. k < Suc i} - {0})" bulwahn@27656: by (auto simp only: insert_Diff) bulwahn@27656: have b: "{k \ M. k < Suc i} - {0} = {k \ M - {0}. k < Suc i}" by auto lp15@57113: from finite_M_bounded_by_nat[of "\x. x \ M" "Suc i"] lp15@57113: have "Suc (card {k. Suc k \ M \ k < i}) = card (insert 0 ({k \ M. k < Suc i} - {0}))" bulwahn@27656: apply (subst card_insert) bulwahn@27656: apply simp_all bulwahn@27656: apply (subst b) bulwahn@27656: apply (subst card_less_Suc2[symmetric]) bulwahn@27656: apply simp_all bulwahn@27656: done bulwahn@27656: with c show ?thesis by simp bulwahn@27656: qed bulwahn@27656: paulson@14485: wenzelm@60758: subsection \Lemmas useful with the summation operator setsum\ paulson@13850: wenzelm@60758: text \For examples, see Algebra/poly/UnivPoly2.thy\ ballarin@13735: wenzelm@60758: subsubsection \Disjoint Unions\ ballarin@13735: wenzelm@60758: text \Singletons and open intervals\ ballarin@13735: ballarin@13735: lemma ivl_disj_un_singleton: nipkow@15045: "{l::'a::linorder} Un {l<..} = {l..}" nipkow@15045: "{.. {l} Un {l<.. {l<.. {l} Un {l<..u} = {l..u}" nipkow@15045: "(l::'a::linorder) <= u ==> {l..One- and two-sided intervals\ ballarin@13735: ballarin@13735: lemma ivl_disj_un_one: nipkow@15045: "(l::'a::linorder) < u ==> {..l} Un {l<.. {.. {..l} Un {l<..u} = {..u}" nipkow@15045: "(l::'a::linorder) <= u ==> {.. {l<..u} Un {u<..} = {l<..}" nipkow@15045: "(l::'a::linorder) < u ==> {l<.. {l..u} Un {u<..} = {l..}" nipkow@15045: "(l::'a::linorder) <= u ==> {l..Two- and two-sided intervals\ ballarin@13735: ballarin@13735: lemma ivl_disj_un_two: nipkow@15045: "[| (l::'a::linorder) < m; m <= u |] ==> {l<.. {l<..m} Un {m<.. {l.. {l..m} Un {m<.. {l<.. {l<..m} Un {m<..u} = {l<..u}" nipkow@15045: "[| (l::'a::linorder) <= m; m <= u |] ==> {l.. {l..m} Un {m<..u} = {l..u}" ballarin@14398: by auto ballarin@13735: lp15@60150: lemma ivl_disj_un_two_touch: lp15@60150: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. {l..m} Un {m.. {l<..m} Un {m..u} = {l<..u}" lp15@60150: "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m..u} = {l..u}" lp15@60150: by auto lp15@60150: lp15@60150: lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch ballarin@13735: wenzelm@60758: subsubsection \Disjoint Intersections\ ballarin@13735: wenzelm@60758: text \One- and two-sided intervals\ ballarin@13735: ballarin@13735: lemma ivl_disj_int_one: nipkow@15045: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ ballarin@13735: ballarin@13735: lemma ivl_disj_int_two: nipkow@15045: "{l::'a::order<..Some Differences\ nipkow@15542: nipkow@15542: lemma ivl_diff[simp]: nipkow@15542: "i \ n \ {i..Some Subset Conditions\ nipkow@15542: blanchet@54147: lemma ivl_subset [simp]: nipkow@15542: "({i.. {m.. i | m \ i & j \ (n::'a::linorder))" nipkow@15542: apply(auto simp:linorder_not_le) nipkow@15542: apply(rule ccontr) nipkow@15542: apply(insert linorder_le_less_linear[of i n]) nipkow@15542: apply(clarsimp simp:linorder_not_le) nipkow@44890: apply(fastforce) nipkow@15542: done nipkow@15542: nipkow@15041: wenzelm@60758: subsection \Summation indexed over intervals\ nipkow@15042: nipkow@15042: syntax nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) nipkow@15048: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) nipkow@16052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) nipkow@15042: syntax (xsymbols) nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) nipkow@15048: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) nipkow@16052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) nipkow@15056: syntax (latex_sum output) nipkow@15052: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" nipkow@15052: ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) nipkow@15052: "_from_upto_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" nipkow@15052: ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10) nipkow@16052: "_upt_setsum" :: "idt \ 'a \ 'b \ 'b" nipkow@16052: ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10) nipkow@15052: "_upto_setsum" :: "idt \ 'a \ 'b \ 'b" nipkow@16052: ("(3\<^raw:$\sum_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) nipkow@15041: nipkow@15048: translations nipkow@28853: "\x=a..b. t" == "CONST setsum (%x. t) {a..b}" nipkow@28853: "\x=a..i\n. t" == "CONST setsum (\i. t) {..n}" nipkow@28853: "\ii. t) {..The above introduces some pretty alternative syntaxes for nipkow@15056: summation over intervals: nipkow@15052: \begin{center} nipkow@15052: \begin{tabular}{lll} nipkow@15056: Old & New & \LaTeX\\ nipkow@15056: @{term[source]"\x\{a..b}. e"} & @{term"\x=a..b. e"} & @{term[mode=latex_sum]"\x=a..b. e"}\\ nipkow@15056: @{term[source]"\x\{a..x=a..x=a..x\{..b}. e"} & @{term"\x\b. e"} & @{term[mode=latex_sum]"\x\b. e"}\\ nipkow@15056: @{term[source]"\x\{..xxlatex_sum\ (e.g.\ via \mode = latex_sum\ in nipkow@15056: antiquotations). It is not the default \LaTeX\ output because it only nipkow@15056: works well with italic-style formulae, not tt-style. nipkow@15052: nipkow@15052: Note that for uniformity on @{typ nat} it is better to use wenzelm@61799: @{term"\x::nat=0..\x: \setsum\ may nipkow@15052: not provide all lemmas available for @{term"{m.. nipkow@15052: wenzelm@60758: text\This congruence rule should be used for sums over intervals as haftmann@57418: the standard theorem @{text[source]setsum.cong} does not work well nipkow@15542: with the simplifier who adds the unsimplified premise @{term"x:B"} to wenzelm@60758: the context.\ nipkow@15542: nipkow@15542: lemma setsum_ivl_cong: nipkow@15542: "\a = c; b = d; !!x. \ c \ x; x < d \ \ f x = g x \ \ nipkow@15542: setsum f {a..i \ Suc n. f i) = (\i \ n. f i) + f(Suc n)" haftmann@57514: by (simp add:atMost_Suc ac_simps) nipkow@16052: nipkow@16041: lemma setsum_lessThan_Suc[simp]: "(\i < Suc n. f i) = (\i < n. f i) + f n" haftmann@57514: by (simp add:lessThan_Suc ac_simps) nipkow@15041: nipkow@15911: lemma setsum_cl_ivl_Suc[simp]: nipkow@15561: "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))" haftmann@57514: by (auto simp:ac_simps atLeastAtMostSuc_conv) nipkow@15561: nipkow@15911: lemma setsum_op_ivl_Suc[simp]: nipkow@15561: "setsum f {m.. nipkow@15561: (\i=n..m+1. f i) = (\i=n..m. f i) + f(m + 1)" haftmann@57514: by (auto simp:ac_simps atLeastAtMostSuc_conv) nipkow@16041: *) nipkow@28068: nipkow@28068: lemma setsum_head: nipkow@28068: fixes n :: nat nipkow@28068: assumes mn: "m <= n" nipkow@28068: shows "(\x\{m..n}. P x) = P m + (\x\{m<..n}. P x)" (is "?lhs = ?rhs") nipkow@28068: proof - nipkow@28068: from mn nipkow@28068: have "{m..n} = {m} \ {m<..n}" nipkow@28068: by (auto intro: ivl_disj_un_singleton) nipkow@28068: hence "?lhs = (\x\{m} \ {m<..n}. P x)" nipkow@28068: by (simp add: atLeast0LessThan) nipkow@28068: also have "\ = ?rhs" by simp nipkow@28068: finally show ?thesis . nipkow@28068: qed nipkow@28068: nipkow@28068: lemma setsum_head_Suc: nipkow@28068: "m \ n \ setsum f {m..n} = f m + setsum f {Suc m..n}" nipkow@28068: by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost) nipkow@28068: nipkow@28068: lemma setsum_head_upt_Suc: nipkow@28068: "m < n \ setsum f {m.. n + 1" nipkow@31501: shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}" nipkow@31501: proof- wenzelm@60758: have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto haftmann@57418: thus ?thesis by (auto simp: ivl_disj_int setsum.union_disjoint nipkow@31501: atLeastSucAtMost_greaterThanAtMost) nipkow@31501: qed nipkow@28068: nipkow@15539: lemma setsum_add_nat_ivl: "\ m \ n; n \ p \ \ nipkow@15539: setsum f {m.. 'a::ab_group_add" nipkow@15539: shows "\ m \ n; n \ p \ \ nipkow@15539: setsum f {m.. ('a::ab_group_add)" nipkow@31505: shows "setsum (\k. f k - f(k + 1)) {(m::nat) .. n} = nipkow@31505: (if m <= n then f m - f(n + 1) else 0)" nipkow@31505: by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) nipkow@31505: hoelzl@56194: lemma setsum_nat_group: "(\m(i,j)\{(i,j). i+j < n}. f i j) = (\ki\k. f i (k - i))" lp15@60150: apply (simp add: setsum.Sigma) lp15@60150: apply (rule setsum.reindex_bij_witness[where j="\(i, j). (i+j, i)" and i="\(k, i). (i, k - i)"]) lp15@60150: apply auto lp15@60150: done lp15@60150: lp15@60150: lemma setsum_triangle_reindex_eq: lp15@60150: fixes n :: nat lp15@60150: shows "(\(i,j)\{(i,j). i+j \ n}. f i j) = (\k\n. \i\k. f i (k - i))" lp15@60150: using setsum_triangle_reindex [of f "Suc n"] lp15@60150: by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lp15@60150: lp15@60162: lemma nat_diff_setsum_reindex: "(\iiShifting bounds\ nipkow@16733: nipkow@15539: lemma setsum_shift_bounds_nat_ivl: nipkow@15539: "setsum f {m+k.. setsum f {Suc 0..k} = setsum f {0..k}" nipkow@28068: by(simp add:setsum_head_Suc) kleing@19106: nipkow@28068: lemma setsum_shift_lb_Suc0_0_upt: nipkow@28068: "f(0::nat) = 0 \ setsum f {Suc 0.. 'a::comm_monoid_add" haftmann@52380: shows "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" haftmann@52380: proof (induct n) haftmann@52380: case 0 show ?case by simp haftmann@52380: next haftmann@52380: case (Suc n) note IH = this haftmann@52380: have "(\i\Suc (Suc n). f i) = (\i\Suc n. f i) + f (Suc (Suc n))" haftmann@52380: by (rule setsum_atMost_Suc) haftmann@52380: also have "(\i\Suc n. f i) = f 0 + (\i\n. f (Suc i))" haftmann@52380: by (rule IH) haftmann@52380: also have "f 0 + (\i\n. f (Suc i)) + f (Suc (Suc n)) = haftmann@52380: f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" haftmann@57512: by (rule add.assoc) haftmann@52380: also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" haftmann@52380: by (rule setsum_atMost_Suc [symmetric]) haftmann@52380: finally show ?case . haftmann@52380: qed haftmann@52380: lp15@56238: lemma setsum_last_plus: fixes n::nat shows "m <= n \ (\i = m..n. f i) = f n + (\i = m.. 'a::ab_group_add" lp15@56238: assumes "m \ Suc n" lp15@56238: shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" lp15@56238: using assms by (induct n) (auto simp: le_Suc_eq) lp15@55718: lp15@55718: lemma nested_setsum_swap: lp15@55718: "(\i = 0..n. (\j = 0..j = 0..i = Suc j..n. a i j)" haftmann@57418: by (induction n) (auto simp: setsum.distrib) lp15@55718: lp15@56215: lemma nested_setsum_swap': lp15@56215: "(\i\n. (\jji = Suc j..n. a i j)" haftmann@57418: by (induction n) (auto simp: setsum.distrib) lp15@56215: lp15@56238: lemma setsum_zero_power' [simp]: lp15@56238: fixes c :: "nat \ 'a::field" lp15@56238: shows "(\i\A. c i * 0^i / d i) = (if finite A \ 0 \ A then c 0 / d 0 else 0)" lp15@56238: using setsum_zero_power [of "\i. c i / d i" A] lp15@56238: by auto lp15@56238: haftmann@52380: eberlm@61524: subsection \Telescoping\ eberlm@61524: eberlm@61524: lemma setsum_telescope: eberlm@61524: fixes f::"nat \ 'a::ab_group_add" eberlm@61524: shows "setsum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" eberlm@61524: by (induct i) simp_all eberlm@61524: eberlm@61524: lemma setsum_telescope'': eberlm@61524: assumes "m \ n" eberlm@61524: shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" eberlm@61524: by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) eberlm@61524: wenzelm@60758: subsection \The formula for geometric sums\ ballarin@17149: ballarin@17149: lemma geometric_sum: haftmann@36307: assumes "x \ 1" hoelzl@56193: shows "(\i 0" by simp_all hoelzl@56193: moreover have "(\iy \ 0\) haftmann@36307: ultimately show ?thesis by simp haftmann@36307: qed haftmann@36307: lp15@60162: lemma diff_power_eq_setsum: lp15@60162: fixes y :: "'a::{comm_ring,monoid_mult}" lp15@60162: shows lp15@60162: "x ^ (Suc n) - y ^ (Suc n) = lp15@60162: (x - y) * (\pppp\\COMPLEX_POLYFUN\ in HOL Light\ lp15@60162: fixes x :: "'a::{comm_ring,monoid_mult}" lp15@60162: shows "x^n - y^n = (x - y) * (\i 0 \ x^n - 1 = (x - 1) * (\i 0 \ 1 - x^n = (1 - x) * (\i 0 \ 1 - x^n = (1 - x) * (\iThe formula for arithmetic sums\ kleing@19469: huffman@47222: lemma gauss_sum: hoelzl@56193: "(2::'a::comm_semiring_1)*(\i\{1..n}. of_nat i) = of_nat n*((of_nat n)+1)" kleing@19469: proof (induct n) kleing@19469: case 0 kleing@19469: show ?case by simp kleing@19469: next kleing@19469: case (Suc n) huffman@47222: then show ?case huffman@47222: by (simp add: algebra_simps add: one_add_one [symmetric] del: one_add_one) huffman@47222: (* FIXME: make numeral cancellation simprocs work for semirings *) kleing@19469: qed kleing@19469: kleing@19469: theorem arith_series_general: huffman@47222: "(2::'a::comm_semiring_1) * (\i\{.. 1" kleing@19469: let ?I = "\i. of_nat i" and ?n = "of_nat n" kleing@19469: have kleing@19469: "(\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)" haftmann@57514: by (simp only: mult.assoc gauss_sum [of "n - 1"], unfold One_nat_def) haftmann@57514: (simp add: mult.commute trans [OF add.commute of_nat_Suc [symmetric]]) huffman@47222: finally show ?thesis huffman@47222: unfolding mult_2 by (simp add: algebra_simps) kleing@19469: next kleing@19469: assume "\(n > 1)" kleing@19469: hence "n = 1 \ n = 0" by auto huffman@47222: thus ?thesis by (auto simp: mult_2) kleing@19469: qed kleing@19469: kleing@19469: lemma arith_series_nat: huffman@47222: "(2::nat) * (\i\{..i\{..i\{..x. Q x \ P x \ (\xxxProducts indexed over intervals\ paulson@29960: paulson@29960: syntax paulson@29960: "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) paulson@29960: "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) paulson@29960: "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) paulson@29960: "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) paulson@29960: syntax (xsymbols) paulson@29960: "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) paulson@29960: "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) paulson@29960: "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) paulson@29960: "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) paulson@29960: syntax (latex_prod output) paulson@29960: "_from_to_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" paulson@29960: ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10) paulson@29960: "_from_upto_setprod" :: "idt \ 'a \ 'a \ 'b \ 'b" paulson@29960: ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10) paulson@29960: "_upt_setprod" :: "idt \ 'a \ 'b \ 'b" paulson@29960: ("(3\<^raw:$\prod_{>_ < _\<^raw:}$> _)" [0,0,10] 10) paulson@29960: "_upto_setprod" :: "idt \ 'a \ 'b \ 'b" paulson@29960: ("(3\<^raw:$\prod_{>_ \ _\<^raw:}$> _)" [0,0,10] 10) paulson@29960: paulson@29960: translations paulson@29960: "\x=a..b. t" == "CONST setprod (%x. t) {a..b}" paulson@29960: "\x=a..i\n. t" == "CONST setprod (\i. t) {..n}" paulson@29960: "\ii. t) {..Transfer setup\ haftmann@33318: haftmann@33318: lemma transfer_nat_int_set_functions: haftmann@33318: "{..n} = nat ` {0..int n}" haftmann@33318: "{m..n} = nat ` {int m..int n}" (* need all variants of these! *) haftmann@33318: apply (auto simp add: image_def) haftmann@33318: apply (rule_tac x = "int x" in bexI) haftmann@33318: apply auto haftmann@33318: apply (rule_tac x = "int x" in bexI) haftmann@33318: apply auto haftmann@33318: done haftmann@33318: haftmann@33318: lemma transfer_nat_int_set_function_closures: haftmann@33318: "x >= 0 \ nat_set {x..y}" haftmann@33318: by (simp add: nat_set_def) haftmann@33318: haftmann@35644: declare transfer_morphism_nat_int[transfer add haftmann@33318: return: transfer_nat_int_set_functions haftmann@33318: transfer_nat_int_set_function_closures haftmann@33318: ] haftmann@33318: haftmann@33318: lemma transfer_int_nat_set_functions: haftmann@33318: "is_nat m \ is_nat n \ {m..n} = int ` {nat m..nat n}" haftmann@33318: by (simp only: is_nat_def transfer_nat_int_set_functions haftmann@33318: transfer_nat_int_set_function_closures haftmann@33318: transfer_nat_int_set_return_embed nat_0_le haftmann@33318: cong: transfer_nat_int_set_cong) haftmann@33318: haftmann@33318: lemma transfer_int_nat_set_function_closures: haftmann@33318: "is_nat x \ nat_set {x..y}" haftmann@33318: by (simp only: transfer_nat_int_set_function_closures is_nat_def) haftmann@33318: haftmann@35644: declare transfer_morphism_int_nat[transfer add haftmann@33318: return: transfer_int_nat_set_functions haftmann@33318: transfer_int_nat_set_function_closures haftmann@33318: ] haftmann@33318: lp15@55242: lemma setprod_int_plus_eq: "setprod int {i..i+j} = \{int i..int (i+j)}" lp15@55242: by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lp15@55242: lp15@55242: lemma setprod_int_eq: "setprod int {i..j} = \{int i..int j}" lp15@55242: proof (cases "i \ j") lp15@55242: case True lp15@55242: then show ?thesis lp15@55242: by (metis Nat.le_iff_add setprod_int_plus_eq) lp15@55242: next lp15@55242: case False lp15@55242: then show ?thesis lp15@55242: by auto lp15@55242: qed lp15@55242: eberlm@61524: eberlm@61524: subsection \Shifting bounds\ eberlm@61524: eberlm@61524: lemma setprod_shift_bounds_nat_ivl: eberlm@61524: "setprod f {m+k.. b \ setprod f {a.. Suc n" eberlm@61524: shows "setprod f {m..Suc n} = f (Suc n) * setprod f {m..n}" eberlm@61524: proof - eberlm@61524: from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto eberlm@61524: also have "setprod f \ = f (Suc n) * setprod f {m..n}" by simp eberlm@61524: finally show ?thesis . eberlm@61524: qed eberlm@61524: nipkow@8924: end