nipkow@8924: (* Title: HOL/SetInterval.thy nipkow@8924: ID: $Id$ ballarin@13735: Author: Tobias Nipkow and Clemens Ballarin paulson@14485: Additions by Jeremy Avigad in March 2004 paulson@8957: Copyright 2000 TU Muenchen nipkow@8924: ballarin@13735: lessThan, greaterThan, atLeast, atMost and two-sided intervals nipkow@8924: *) nipkow@8924: wenzelm@14577: header {* Set intervals *} wenzelm@14577: paulson@14485: theory SetInterval = IntArith: nipkow@8924: nipkow@8924: constdefs wenzelm@11609: lessThan :: "('a::ord) => 'a set" ("(1{.._'(})") wenzelm@11609: "{..u(} == {x. x 'a set" ("(1{.._})") wenzelm@11609: "{..u} == {x. x<=u}" nipkow@8924: wenzelm@11609: greaterThan :: "('a::ord) => 'a set" ("(1{')_..})") wenzelm@11609: "{)l..} == {x. l 'a set" ("(1{_..})") wenzelm@11609: "{l..} == {x. l<=x}" nipkow@8924: ballarin@13735: greaterThanLessThan :: "['a::ord, 'a] => 'a set" ("(1{')_.._'(})") ballarin@13735: "{)l..u(} == {)l..} Int {..u(}" ballarin@13735: ballarin@13735: atLeastLessThan :: "['a::ord, 'a] => 'a set" ("(1{_.._'(})") ballarin@13735: "{l..u(} == {l..} Int {..u(}" ballarin@13735: ballarin@13735: greaterThanAtMost :: "['a::ord, 'a] => 'a set" ("(1{')_.._})") ballarin@13735: "{)l..u} == {)l..} Int {..u}" ballarin@13735: ballarin@13735: atLeastAtMost :: "['a::ord, 'a] => 'a set" ("(1{_.._})") ballarin@13735: "{l..u} == {l..} Int {..u}" ballarin@13735: kleing@14418: syntax kleing@14418: "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3UN _<=_./ _)" 10) kleing@14418: "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3UN _<_./ _)" 10) kleing@14418: "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3INT _<=_./ _)" 10) kleing@14418: "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3INT _<_./ _)" 10) kleing@14418: kleing@14418: syntax (input) kleing@14418: "@UNION_le" :: "nat => nat => 'b set => 'b set" ("(3\ _\_./ _)" 10) kleing@14418: "@UNION_less" :: "nat => nat => 'b set => 'b set" ("(3\ _<_./ _)" 10) kleing@14418: "@INTER_le" :: "nat => nat => 'b set => 'b set" ("(3\ _\_./ _)" 10) kleing@14418: "@INTER_less" :: "nat => nat => 'b set => 'b set" ("(3\ _<_./ _)" 10) kleing@14418: kleing@14418: syntax (xsymbols) wenzelm@14846: "@UNION_le" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ \ _\<^esub>)/ _)" 10) wenzelm@14846: "@UNION_less" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ < _\<^esub>)/ _)" 10) wenzelm@14846: "@INTER_le" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ \ _\<^esub>)/ _)" 10) wenzelm@14846: "@INTER_less" :: "nat \ nat => 'b set => 'b set" ("(3\(00\<^bsub>_ < _\<^esub>)/ _)" 10) kleing@14418: kleing@14418: translations kleing@14418: "UN i<=n. A" == "UN i:{..n}. A" kleing@14418: "UN i atLeast y) = (y \ (x::'a::order))" paulson@13850: by (blast intro: order_trans) paulson@13850: paulson@13850: lemma atLeast_eq_iff [iff]: paulson@13850: "(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@13850: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" paulson@13850: apply (auto simp add: greaterThan_def) paulson@13850: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma greaterThan_eq_iff [iff]: paulson@13850: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" paulson@13850: apply (rule iffI) paulson@13850: apply (erule equalityE) paulson@13850: apply (simp add: greaterThan_subset_iff order_antisym, simp) paulson@13850: done paulson@13850: paulson@13850: lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::order))" paulson@13850: by (blast intro: order_trans) paulson@13850: paulson@13850: 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@13850: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" paulson@13850: apply (auto simp add: lessThan_def) paulson@13850: apply (subst linorder_not_less [symmetric], blast) paulson@13850: done paulson@13850: paulson@13850: lemma lessThan_eq_iff [iff]: paulson@13850: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" paulson@13850: apply (rule iffI) paulson@13850: apply (erule equalityE) paulson@13850: apply (simp add: lessThan_subset_iff order_antisym, simp) ballarin@13735: done ballarin@13735: ballarin@13735: paulson@13850: subsection {*Two-sided intervals*} ballarin@13735: wenzelm@14577: text {* @{text greaterThanLessThan} *} ballarin@13735: ballarin@13735: lemma greaterThanLessThan_iff [simp]: ballarin@13735: "(i : {)l..u(}) = (l < i & i < u)" ballarin@13735: by (simp add: greaterThanLessThan_def) ballarin@13735: wenzelm@14577: text {* @{text atLeastLessThan} *} ballarin@13735: ballarin@13735: lemma atLeastLessThan_iff [simp]: ballarin@13735: "(i : {l..u(}) = (l <= i & i < u)" ballarin@13735: by (simp add: atLeastLessThan_def) ballarin@13735: wenzelm@14577: text {* @{text greaterThanAtMost} *} ballarin@13735: ballarin@13735: lemma greaterThanAtMost_iff [simp]: ballarin@13735: "(i : {)l..u}) = (l < i & i <= u)" ballarin@13735: by (simp add: greaterThanAtMost_def) ballarin@13735: wenzelm@14577: text {* @{text atLeastAtMost} *} ballarin@13735: ballarin@13735: lemma atLeastAtMost_iff [simp]: ballarin@13735: "(i : {l..u}) = (l <= i & i <= u)" ballarin@13735: by (simp add: atLeastAtMost_def) ballarin@13735: wenzelm@14577: text {* The above four lemmas could be declared as iffs. wenzelm@14577: If we do so, a call to blast in Hyperreal/Star.ML, lemma @{text STAR_Int} wenzelm@14577: seems to take forever (more than one hour). *} ballarin@13735: paulson@14485: paulson@14485: subsection {* Intervals of natural numbers *} paulson@14485: 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: paulson@14485: lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" paulson@14485: by (simp add: lessThan_def atMost_def less_Suc_eq_le) paulson@14485: paulson@14485: lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV" paulson@14485: by blast paulson@14485: 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: 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: 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: nipkow@15042: lemma atLeast0LessThan [simp]: "{0::nat..n(} = {..n(}" nipkow@15042: by(simp add:lessThan_def atLeastLessThan_def) nipkow@15042: wenzelm@14577: text {* Intervals of nats with @{text Suc} *} paulson@14485: paulson@14485: lemma atLeastLessThanSuc_atLeastAtMost: "{l..Suc u(} = {l..u}" paulson@14485: by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def) paulson@14485: paulson@14485: lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {)l..u}" paulson@14485: by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def paulson@14485: greaterThanAtMost_def) paulson@14485: paulson@14485: lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..u(} = {)l..u(}" paulson@14485: by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def paulson@14485: greaterThanLessThan_def) paulson@14485: paulson@14485: subsubsection {* Finiteness *} paulson@14485: paulson@14485: lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..k(}" paulson@14485: by (induct k) (simp_all add: lessThan_Suc) paulson@14485: paulson@14485: lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}" paulson@14485: by (induct k) (simp_all add: atMost_Suc) paulson@14485: paulson@14485: lemma finite_greaterThanLessThan [iff]: paulson@14485: fixes l :: nat shows "finite {)l..u(}" paulson@14485: by (simp add: greaterThanLessThan_def) paulson@14485: paulson@14485: lemma finite_atLeastLessThan [iff]: paulson@14485: fixes l :: nat shows "finite {l..u(}" paulson@14485: by (simp add: atLeastLessThan_def) paulson@14485: paulson@14485: lemma finite_greaterThanAtMost [iff]: paulson@14485: fixes l :: nat shows "finite {)l..u}" paulson@14485: by (simp add: greaterThanAtMost_def) paulson@14485: paulson@14485: lemma finite_atLeastAtMost [iff]: paulson@14485: fixes l :: nat shows "finite {l..u}" paulson@14485: by (simp add: atLeastAtMost_def) paulson@14485: paulson@14485: lemma bounded_nat_set_is_finite: paulson@14485: "(ALL i:N. i < (n::nat)) ==> finite N" paulson@14485: -- {* A bounded set of natural numbers is finite. *} paulson@14485: apply (rule finite_subset) paulson@14485: apply (rule_tac [2] finite_lessThan, auto) paulson@14485: done paulson@14485: paulson@14485: subsubsection {* Cardinality *} paulson@14485: paulson@14485: lemma card_lessThan [simp]: "card {..u(} = u" paulson@14485: by (induct_tac u, simp_all add: lessThan_Suc) paulson@14485: paulson@14485: lemma card_atMost [simp]: "card {..u} = Suc u" paulson@14485: by (simp add: lessThan_Suc_atMost [THEN sym]) paulson@14485: paulson@14485: lemma card_atLeastLessThan [simp]: "card {l..u(} = u - l" paulson@14485: apply (subgoal_tac "card {l..u(} = card {..u-l(}") paulson@14485: apply (erule ssubst, rule card_lessThan) paulson@14485: apply (subgoal_tac "(%x. x + l) ` {..u-l(} = {l..u(}") paulson@14485: apply (erule subst) paulson@14485: apply (rule card_image) paulson@14485: apply (rule finite_lessThan) paulson@14485: apply (simp add: inj_on_def) paulson@14485: apply (auto simp add: image_def atLeastLessThan_def lessThan_def) paulson@14485: apply arith paulson@14485: apply (rule_tac x = "x - l" in exI) paulson@14485: apply arith paulson@14485: done paulson@14485: paulson@14485: lemma card_atLeastAtMost [simp]: "card {l..u} = Suc u - l" paulson@14485: by (subst atLeastLessThanSuc_atLeastAtMost [THEN sym], simp) paulson@14485: paulson@14485: lemma card_greaterThanAtMost [simp]: "card {)l..u} = u - l" paulson@14485: by (subst atLeastSucAtMost_greaterThanAtMost [THEN sym], simp) paulson@14485: paulson@14485: lemma card_greaterThanLessThan [simp]: "card {)l..u(} = u - Suc l" paulson@14485: by (subst atLeastSucLessThan_greaterThanLessThan [THEN sym], simp) paulson@14485: paulson@14485: subsection {* Intervals of integers *} paulson@14485: paulson@14485: lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..u+1(} = {l..(u::int)}" paulson@14485: by (auto simp add: atLeastAtMost_def atLeastLessThan_def) paulson@14485: paulson@14485: lemma atLeastPlusOneAtMost_greaterThanAtMost_int: "{l+1..u} = {)l..(u::int)}" paulson@14485: by (auto simp add: atLeastAtMost_def greaterThanAtMost_def) paulson@14485: paulson@14485: lemma atLeastPlusOneLessThan_greaterThanLessThan_int: paulson@14485: "{l+1..u(} = {)l..(u::int)(}" paulson@14485: by (auto simp add: atLeastLessThan_def greaterThanLessThan_def) paulson@14485: paulson@14485: subsubsection {* Finiteness *} paulson@14485: paulson@14485: lemma image_atLeastZeroLessThan_int: "0 \ u ==> paulson@14485: {(0::int)..u(} = int ` {..nat u(}" paulson@14485: apply (unfold image_def lessThan_def) paulson@14485: apply auto paulson@14485: apply (rule_tac x = "nat x" in exI) paulson@14485: apply (auto simp add: zless_nat_conj zless_nat_eq_int_zless [THEN sym]) paulson@14485: done paulson@14485: paulson@14485: lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..u(}" paulson@14485: apply (case_tac "0 \ u") paulson@14485: apply (subst image_atLeastZeroLessThan_int, assumption) paulson@14485: apply (rule finite_imageI) paulson@14485: apply auto paulson@14485: apply (subgoal_tac "{0..u(} = {}") paulson@14485: apply auto paulson@14485: done paulson@14485: paulson@14485: lemma image_atLeastLessThan_int_shift: paulson@14485: "(%x. x + (l::int)) ` {0..u-l(} = {l..u(}" paulson@14485: apply (auto simp add: image_def atLeastLessThan_iff) paulson@14485: apply (rule_tac x = "x - l" in bexI) paulson@14485: apply auto paulson@14485: done paulson@14485: paulson@14485: lemma finite_atLeastLessThan_int [iff]: "finite {l..(u::int)(}" paulson@14485: apply (subgoal_tac "(%x. x + l) ` {0..u-l(} = {l..u(}") paulson@14485: apply (erule subst) paulson@14485: apply (rule finite_imageI) paulson@14485: apply (rule finite_atLeastZeroLessThan_int) paulson@14485: apply (rule image_atLeastLessThan_int_shift) paulson@14485: done paulson@14485: paulson@14485: lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}" paulson@14485: by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp) paulson@14485: paulson@14485: lemma finite_greaterThanAtMost_int [iff]: "finite {)l..(u::int)}" paulson@14485: by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp) paulson@14485: paulson@14485: lemma finite_greaterThanLessThan_int [iff]: "finite {)l..(u::int)(}" paulson@14485: by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp) paulson@14485: paulson@14485: subsubsection {* Cardinality *} paulson@14485: paulson@14485: lemma card_atLeastZeroLessThan_int: "card {(0::int)..u(} = nat u" paulson@14485: apply (case_tac "0 \ 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: paulson@14485: lemma card_atLeastLessThan_int [simp]: "card {l..u(} = nat (u - l)" paulson@14485: apply (subgoal_tac "card {l..u(} = card {0..u-l(}") paulson@14485: apply (erule ssubst, rule card_atLeastZeroLessThan_int) paulson@14485: apply (subgoal_tac "(%x. x + l) ` {0..u-l(} = {l..u(}") paulson@14485: apply (erule subst) paulson@14485: apply (rule card_image) paulson@14485: apply (rule finite_atLeastZeroLessThan_int) paulson@14485: apply (simp add: inj_on_def) paulson@14485: apply (rule image_atLeastLessThan_int_shift) paulson@14485: done paulson@14485: paulson@14485: lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)" paulson@14485: apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym]) paulson@14485: apply (auto simp add: compare_rls) paulson@14485: done paulson@14485: paulson@14485: lemma card_greaterThanAtMost_int [simp]: "card {)l..u} = nat (u - l)" paulson@14485: by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp) paulson@14485: paulson@14485: lemma card_greaterThanLessThan_int [simp]: "card {)l..u(} = nat (u - (l + 1))" paulson@14485: by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp) paulson@14485: paulson@14485: paulson@13850: subsection {*Lemmas useful with the summation operator setsum*} paulson@13850: wenzelm@14577: text {* For examples, see Algebra/poly/UnivPoly.thy *} ballarin@13735: wenzelm@14577: subsubsection {* Disjoint Unions *} ballarin@13735: wenzelm@14577: text {* Singletons and open intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_un_singleton: ballarin@13735: "{l::'a::linorder} Un {)l..} = {l..}" ballarin@13735: "{..u(} Un {u::'a::linorder} = {..u}" ballarin@13735: "(l::'a::linorder) < u ==> {l} Un {)l..u(} = {l..u(}" ballarin@13735: "(l::'a::linorder) < u ==> {)l..u(} Un {u} = {)l..u}" ballarin@13735: "(l::'a::linorder) <= u ==> {l} Un {)l..u} = {l..u}" ballarin@13735: "(l::'a::linorder) <= u ==> {l..u(} Un {u} = {l..u}" ballarin@14398: by auto ballarin@13735: wenzelm@14577: text {* One- and two-sided intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_un_one: ballarin@13735: "(l::'a::linorder) < u ==> {..l} Un {)l..u(} = {..u(}" ballarin@13735: "(l::'a::linorder) <= u ==> {..l(} Un {l..u(} = {..u(}" ballarin@13735: "(l::'a::linorder) <= u ==> {..l} Un {)l..u} = {..u}" ballarin@13735: "(l::'a::linorder) <= u ==> {..l(} Un {l..u} = {..u}" ballarin@13735: "(l::'a::linorder) <= u ==> {)l..u} Un {)u..} = {)l..}" ballarin@13735: "(l::'a::linorder) < u ==> {)l..u(} Un {u..} = {)l..}" ballarin@13735: "(l::'a::linorder) <= u ==> {l..u} Un {)u..} = {l..}" ballarin@13735: "(l::'a::linorder) <= u ==> {l..u(} Un {u..} = {l..}" ballarin@14398: by auto ballarin@13735: wenzelm@14577: text {* Two- and two-sided intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_un_two: ballarin@13735: "[| (l::'a::linorder) < m; m <= u |] ==> {)l..m(} Un {m..u(} = {)l..u(}" ballarin@13735: "[| (l::'a::linorder) <= m; m < u |] ==> {)l..m} Un {)m..u(} = {)l..u(}" ballarin@13735: "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m(} Un {m..u(} = {l..u(}" ballarin@13735: "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {)m..u(} = {l..u(}" ballarin@13735: "[| (l::'a::linorder) < m; m <= u |] ==> {)l..m(} Un {m..u} = {)l..u}" ballarin@13735: "[| (l::'a::linorder) <= m; m <= u |] ==> {)l..m} Un {)m..u} = {)l..u}" ballarin@13735: "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m(} Un {m..u} = {l..u}" ballarin@13735: "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {)m..u} = {l..u}" ballarin@14398: by auto ballarin@13735: ballarin@13735: lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ballarin@13735: wenzelm@14577: subsubsection {* Disjoint Intersections *} ballarin@13735: wenzelm@14577: text {* Singletons and open intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_int_singleton: ballarin@13735: "{l::'a::order} Int {)l..} = {}" ballarin@13735: "{..u(} Int {u} = {}" ballarin@13735: "{l} Int {)l..u(} = {}" ballarin@13735: "{)l..u(} Int {u} = {}" ballarin@13735: "{l} Int {)l..u} = {}" ballarin@13735: "{l..u(} Int {u} = {}" ballarin@13735: by simp+ ballarin@13735: wenzelm@14577: text {* One- and two-sided intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_int_one: ballarin@13735: "{..l::'a::order} Int {)l..u(} = {}" ballarin@13735: "{..l(} Int {l..u(} = {}" ballarin@13735: "{..l} Int {)l..u} = {}" ballarin@13735: "{..l(} Int {l..u} = {}" ballarin@13735: "{)l..u} Int {)u..} = {}" ballarin@13735: "{)l..u(} Int {u..} = {}" ballarin@13735: "{l..u} Int {)u..} = {}" ballarin@13735: "{l..u(} Int {u..} = {}" ballarin@14398: by auto ballarin@13735: wenzelm@14577: text {* Two- and two-sided intervals *} ballarin@13735: ballarin@13735: lemma ivl_disj_int_two: ballarin@13735: "{)l::'a::order..m(} Int {m..u(} = {}" ballarin@13735: "{)l..m} Int {)m..u(} = {}" ballarin@13735: "{l..m(} Int {m..u(} = {}" ballarin@13735: "{l..m} Int {)m..u(} = {}" ballarin@13735: "{)l..m(} Int {m..u} = {}" ballarin@13735: "{)l..m} Int {)m..u} = {}" ballarin@13735: "{l..m(} Int {m..u} = {}" ballarin@13735: "{l..m} Int {)m..u} = {}" ballarin@14398: by auto ballarin@13735: ballarin@13735: lemmas ivl_disj_int = ivl_disj_int_singleton ivl_disj_int_one ivl_disj_int_two ballarin@13735: nipkow@15041: nipkow@15042: subsection {* Summation indexed over intervals *} nipkow@15042: nipkow@15042: text{* We introduce the obvious syntax @{text"\x=a..b. e"} for nipkow@15042: @{term"\x\{a..b}. e"}. *} nipkow@15042: nipkow@15042: syntax nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,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@15042: syntax (HTML output) nipkow@15042: "_from_to_setsum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) nipkow@15042: nipkow@15042: translations "\x=a..b. t" == "setsum (%x. t) {a..b}" nipkow@15042: nipkow@15042: nipkow@15042: subsection {* Summation up to *} nipkow@15041: nipkow@15041: text{* Legacy, only used in HoareParallel and Isar-Examples. Really nipkow@15042: needed? Probably better to replace it with above syntax. *} nipkow@15041: nipkow@15041: syntax nipkow@15042: "_Summation" :: "idt => 'a => 'b => 'b" ("\_<_. _" [0, 51, 10] 10) nipkow@15041: translations nipkow@15042: "\i < n. b" == "setsum (\i. b) {..n(}" nipkow@15041: nipkow@15041: lemma Summation_Suc[simp]: "(\i < Suc n. b i) = b n + (\i < n. b i)" nipkow@15041: by (simp add:lessThan_Suc) nipkow@15041: nipkow@8924: end