diff -r cedf3480fdad -r 3db8520941a4 src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Wed Jul 11 23:24:25 2018 +0100 +++ b/src/HOL/Set_Interval.thy Thu Jul 12 17:22:39 2018 +0100 @@ -91,29 +91,27 @@ lemma Compl_lessThan [simp]: "!!k:: 'a::linorder. -lessThan k = atLeast k" -apply (auto simp add: lessThan_def atLeast_def) -done + by (auto simp add: lessThan_def atLeast_def) lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}" -by auto + by auto lemma (in ord) greaterThan_iff [iff]: "(i \ greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) -lemma Compl_atLeast [simp]: - "!!k:: 'a::linorder. -atLeast k = lessThan k" +lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" @@ -146,35 +144,25 @@ 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 + unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" -apply (rule iffI) - apply (erule equalityE) - apply simp_all -done + by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::order))" -by (blast intro: order_trans) + 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) + 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 + unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" -apply (rule iffI) - apply (erule equalityE) - apply simp_all -done + by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" @@ -195,21 +183,17 @@ context ord begin -lemma greaterThanLessThan_iff [simp]: - "(i \ {l<.. i < u)" -by (simp add: greaterThanLessThan_def) - -lemma atLeastLessThan_iff [simp]: - "(i \ {l.. i \ i < u)" -by (simp add: atLeastLessThan_def) - -lemma greaterThanAtMost_iff [simp]: - "(i \ {l<..u}) = (l < i \ i \ u)" -by (simp add: greaterThanAtMost_def) - -lemma atLeastAtMost_iff [simp]: - "(i \ {l..u}) = (l \ i \ i \ u)" -by (simp add: atLeastAtMost_def) +lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" + by (simp add: greaterThanLessThan_def) + +lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" + by (simp add: atLeastLessThan_def) + +lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" + by (simp add: greaterThanAtMost_def) + +lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" + by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them @@ -467,11 +451,11 @@ lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. 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 + apply (auto simp:subset_eq Ball_def not_le) + apply(frule_tac x=a in spec) + apply(erule_tac x=d in allE) + apply (auto simp: ) + done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" @@ -718,17 +702,15 @@ subsubsection \The Constant @{term greaterThan}\ lemma greaterThan_0: "greaterThan 0 = range Suc" -apply (simp add: greaterThan_def) -apply (blast dest: gr0_conv_Suc [THEN iffD1]) -done + unfolding greaterThan_def + by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" -apply (simp add: greaterThan_def) -apply (auto elim: linorder_neqE) -done + unfolding greaterThan_def + by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}" -by blast + by blast subsubsection \The Constant @{term atLeast}\ @@ -736,29 +718,24 @@ 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 + unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) 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 + by blast subsubsection \The Constant @{term atMost}\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" -by (simp add: atMost_def) + 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 + unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV" -by blast + by blast subsubsection \The Constant @{term atLeastLessThan}\ @@ -770,28 +747,24 @@ specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat.. j \ {i.. {j..Intervals and numerals\ @@ -1123,7 +1088,7 @@ by auto lemma image_add_int_atLeastLessThan: - "(%x. x + (l::int)) ` {0..x. x + (l::int)) ` {0..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" -apply (rule finite_subset) - apply (rule_tac [2] finite_lessThan, auto) -done + by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: @@ -1195,11 +1158,9 @@ assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed -lemma finite_nat_set_iff_bounded_le: - "finite(N::nat set) = (\m. \n\N. n<=m)" -apply(simp add:finite_nat_set_iff_bounded) -apply(blast dest:less_imp_le_nat le_imp_less_Suc) -done +lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" + unfolding finite_nat_set_iff_bounded + by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" @@ -1298,10 +1259,9 @@ 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]) + apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) + apply auto + apply (force simp add: atLeastLessThan_add_Un [of 0])+ done @@ -1315,7 +1275,7 @@ lemma card_atLeastLessThan [simp]: "card {l..x. x + l) ` {.. \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \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 + 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 + by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" @@ -1428,26 +1388,21 @@ lemma image_atLeastZeroLessThan_int: "0 \ u ==> {(0::int).. u") - apply (subst image_atLeastZeroLessThan_int, assumption) - apply (rule finite_imageI) - apply auto - done +proof (cases "0 \ u") + case True + then show ?thesis + by (auto simp: image_atLeastZeroLessThan_int) +qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") - apply (subst image_atLeastZeroLessThan_int, assumption) - apply (subst card_image) - apply (auto simp add: inj_on_def) - done +proof (cases "0 \ u") + case True + then show ?thesis + by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) +qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - @@ -1496,43 +1452,44 @@ qed lemma card_less: -assumes zero_in_M: "0 \ M" -shows "card {k \ M. k < Suc i} \ 0" + assumes zero_in_M: "0 \ 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 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_Suc2: + assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" +proof - + have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j + by (cases j) (use assms in auto) + show ?thesis + proof (rule card_bij_eq) + show "inj_on Suc {k. Suc k \ M \ k < i}" + by force + show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" + by (rule inj_on_diff_nat) (use * in blast) + qed (use * in auto) +qed lemma card_less_Suc: - assumes zero_in_M: "0 \ M" + assumes "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 + have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" + by simp + also have "\ = Suc (card {k \ M - {0}. k < Suc i})" + apply (subst card_less_Suc2) + using assms by auto + also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" + by (force intro: arg_cong [where f=card]) + also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" + by (simp add: card_insert) + also have "... = card {k \ M. k < Suc i}" + using assms + by (force simp add: intro: arg_cong [where f=card]) + finally show ?thesis. qed @@ -1549,41 +1506,41 @@ "{.. {l} Un {l<.. {l<.. {l} Un {l<..u} = {l..u}" - "(l::'a::linorder) <= u ==> {l.. u ==> {l} Un {l<..u} = {l..u}" + "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. {.. {..l} Un {l<..u} = {..u}" - "(l::'a::linorder) <= u ==> {.. {l<..u} Un {u<..} = {l<..}" + "(l::'a::linorder) \ u ==> {.. u ==> {..l} Un {l<..u} = {..u}" + "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. {l..u} Un {u<..} = {l..}" - "(l::'a::linorder) <= u ==> {l.. u ==> {l..u} Un {u<..} = {l..}" + "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: - "[| (l::'a::linorder) < m; m <= u |] ==> {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}" + "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" + "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. {l..m} Un {m.. {l<..m} Un {m..u} = {l<..u}" - "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m..u} = {l..u}" + "[| (l::'a::linorder) \ m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" + "[| (l::'a::linorder) \ m; m \ u |] ==> {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 ivl_disj_un_two_touch @@ -1635,14 +1592,11 @@ subsubsection \Some Subset Conditions\ -lemma ivl_subset [simp]: - "({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 +lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" + using linorder_class.le_less_linear[of i n] + apply (auto simp: linorder_not_le) + apply (force intro: leI)+ + done subsection \Generic big monoid operation over intervals\ @@ -1866,14 +1820,14 @@ "sum f {m.. +lemma sum_cl_ivl_add_one_nat: "(n::nat) \ m + 1 ==> (\i=n..m+1. f i) = (\i=n..m. f i) + f(m + 1)" by (auto simp:ac_simps atLeastAtMostSuc_conv) *) lemma sum_head: fixes n :: nat - assumes mn: "m <= n" + assumes mn: "m \ n" shows "(\x\{m..n}. P x) = P m + (\x\{m<..n}. P x)" (is "?lhs = ?rhs") proof - from mn @@ -1904,24 +1858,25 @@ lemmas sum_add_nat_ivl = sum.atLeastLessThan_concat lemma sum_diff_nat_ivl: -fixes f :: "nat \ 'a::ab_group_add" -shows "\ m \ n; n \ p \ \ - sum f {m.. 'a::ab_group_add" + shows "\ m \ n; n \ p \ \ sum f {m.. ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = - (if m <= n then f m - f(n + 1) else 0)" + (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 sum_nat_group: "(\m 0 < k", auto) - apply (induct "n") - apply (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric]) - done +proof (cases k) + case (Suc l) + then have "k > 0" + by auto + then show ?thesis + by (induct n) (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric]) +qed auto lemma sum_triangle_reindex: fixes n :: nat @@ -1971,19 +1926,19 @@ subsubsection \Shifting bounds\ lemma sum_shift_bounds_nat_ivl: - "sum f {m+k..i. f(i + k)){m..i. f(i + k)){m..n::nat}" by (rule sum.reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary sum_shift_bounds_cl_Suc_ivl: - "sum f {Suc m..Suc n} = sum (%i. f(Suc i)){m..n}" + "sum f {Suc m..Suc n} = sum (\i. f(Suc i)){m..n}" by (simp add:sum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary sum_shift_bounds_Suc_ivl: - "sum f {Suc m..i. f(Suc i)){m..f 0 = 0\ by simp qed -lemma sum_shift_lb_Suc0_0: - "sum f {Suc 0..k} = sum f {0..k}" +lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis @@ -2054,7 +2008,7 @@ shows "(\i\n. f i) = f 0 + (\i (\i = m..n. f i) = f n + (\i = m.. n \ (\i = m..n. f i) = f n + (\i = m..Shifting bounds\ lemma prod_shift_bounds_nat_ivl: - "prod f {m+k..i. f(i + k)){m..i. f(i + k)){m..n::nat}" by (rule prod.reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary prod_shift_bounds_cl_Suc_ivl: - "prod f {Suc m..Suc n} = prod (%i. f(Suc i)){m..n}" + "prod f {Suc m..Suc n} = prod (\i. f(Suc i)){m..n}" by (simp add:prod_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary prod_shift_bounds_Suc_ivl: - "prod f {Suc m..i. f(Suc i)){m..