src/HOL/SetInterval.thy
author nipkow
Fri, 16 Jul 2004 11:46:59 +0200
changeset 15052 cc562a263609
parent 15048 11b4dce71d73
child 15056 b75073d90bff
permissions -rw-r--r--
Added nice latex syntax.
     1 (*  Title:      HOL/SetInterval.thy
     2     ID:         $Id$
     3     Author:     Tobias Nipkow and Clemens Ballarin
     4                 Additions by Jeremy Avigad in March 2004
     5     Copyright   2000  TU Muenchen
     6 
     7 lessThan, greaterThan, atLeast, atMost and two-sided intervals
     8 *)
     9 
    10 header {* Set intervals *}
    11 
    12 theory SetInterval = IntArith:
    13 
    14 constdefs
    15   lessThan    :: "('a::ord) => 'a set"	("(1{..<_})")
    16   "{..<u} == {x. x<u}"
    17 
    18   atMost      :: "('a::ord) => 'a set"	("(1{.._})")
    19   "{..u} == {x. x<=u}"
    20 
    21   greaterThan :: "('a::ord) => 'a set"	("(1{_<..})")
    22   "{l<..} == {x. l<x}"
    23 
    24   atLeast     :: "('a::ord) => 'a set"	("(1{_..})")
    25   "{l..} == {x. l<=x}"
    26 
    27   greaterThanLessThan :: "['a::ord, 'a] => 'a set"  ("(1{_<..<_})")
    28   "{l<..<u} == {l<..} Int {..<u}"
    29 
    30   atLeastLessThan :: "['a::ord, 'a] => 'a set"      ("(1{_..<_})")
    31   "{l..<u} == {l..} Int {..<u}"
    32 
    33   greaterThanAtMost :: "['a::ord, 'a] => 'a set"    ("(1{_<.._})")
    34   "{l<..u} == {l<..} Int {..u}"
    35 
    36   atLeastAtMost :: "['a::ord, 'a] => 'a set"        ("(1{_.._})")
    37   "{l..u} == {l..} Int {..u}"
    38 
    39 (* Old syntax, will disappear! *)
    40 syntax
    41   "_lessThan"    :: "('a::ord) => 'a set"	("(1{.._'(})")
    42   "_greaterThan" :: "('a::ord) => 'a set"	("(1{')_..})")
    43   "_greaterThanLessThan" :: "['a::ord, 'a] => 'a set"  ("(1{')_.._'(})")
    44   "_atLeastLessThan" :: "['a::ord, 'a] => 'a set"      ("(1{_.._'(})")
    45   "_greaterThanAtMost" :: "['a::ord, 'a] => 'a set"    ("(1{')_.._})")
    46 translations
    47   "{..m(}" => "{..<m}"
    48   "{)m..}" => "{m<..}"
    49   "{)m..n(}" => "{m<..<n}"
    50   "{m..n(}" => "{m..<n}"
    51   "{)m..n}" => "{m<..n}"
    52 
    53 
    54 text{* A note of warning when using @{term"{..<n}"} on type @{typ
    55 nat}: it is equivalent to @{term"{0::nat..<n}"} but some lemmas involving
    56 @{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *}
    57 
    58 syntax
    59   "@UNION_le"   :: "nat => nat => 'b set => 'b set"       ("(3UN _<=_./ _)" 10)
    60   "@UNION_less" :: "nat => nat => 'b set => 'b set"       ("(3UN _<_./ _)" 10)
    61   "@INTER_le"   :: "nat => nat => 'b set => 'b set"       ("(3INT _<=_./ _)" 10)
    62   "@INTER_less" :: "nat => nat => 'b set => 'b set"       ("(3INT _<_./ _)" 10)
    63 
    64 syntax (input)
    65   "@UNION_le"   :: "nat => nat => 'b set => 'b set"       ("(3\<Union> _\<le>_./ _)" 10)
    66   "@UNION_less" :: "nat => nat => 'b set => 'b set"       ("(3\<Union> _<_./ _)" 10)
    67   "@INTER_le"   :: "nat => nat => 'b set => 'b set"       ("(3\<Inter> _\<le>_./ _)" 10)
    68   "@INTER_less" :: "nat => nat => 'b set => 'b set"       ("(3\<Inter> _<_./ _)" 10)
    69 
    70 syntax (xsymbols)
    71   "@UNION_le"   :: "nat \<Rightarrow> nat => 'b set => 'b set"       ("(3\<Union>(00\<^bsub>_ \<le> _\<^esub>)/ _)" 10)
    72   "@UNION_less" :: "nat \<Rightarrow> nat => 'b set => 'b set"       ("(3\<Union>(00\<^bsub>_ < _\<^esub>)/ _)" 10)
    73   "@INTER_le"   :: "nat \<Rightarrow> nat => 'b set => 'b set"       ("(3\<Inter>(00\<^bsub>_ \<le> _\<^esub>)/ _)" 10)
    74   "@INTER_less" :: "nat \<Rightarrow> nat => 'b set => 'b set"       ("(3\<Inter>(00\<^bsub>_ < _\<^esub>)/ _)" 10)
    75 
    76 translations
    77   "UN i<=n. A"  == "UN i:{..n}. A"
    78   "UN i<n. A"   == "UN i:{..<n}. A"
    79   "INT i<=n. A" == "INT i:{..n}. A"
    80   "INT i<n. A"  == "INT i:{..<n}. A"
    81 
    82 
    83 subsection {* Various equivalences *}
    84 
    85 lemma lessThan_iff [iff]: "(i: lessThan k) = (i<k)"
    86 by (simp add: lessThan_def)
    87 
    88 lemma Compl_lessThan [simp]: 
    89     "!!k:: 'a::linorder. -lessThan k = atLeast k"
    90 apply (auto simp add: lessThan_def atLeast_def)
    91 done
    92 
    93 lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}"
    94 by auto
    95 
    96 lemma greaterThan_iff [iff]: "(i: greaterThan k) = (k<i)"
    97 by (simp add: greaterThan_def)
    98 
    99 lemma Compl_greaterThan [simp]: 
   100     "!!k:: 'a::linorder. -greaterThan k = atMost k"
   101 apply (simp add: greaterThan_def atMost_def le_def, auto)
   102 done
   103 
   104 lemma Compl_atMost [simp]: "!!k:: 'a::linorder. -atMost k = greaterThan k"
   105 apply (subst Compl_greaterThan [symmetric])
   106 apply (rule double_complement) 
   107 done
   108 
   109 lemma atLeast_iff [iff]: "(i: atLeast k) = (k<=i)"
   110 by (simp add: atLeast_def)
   111 
   112 lemma Compl_atLeast [simp]: 
   113     "!!k:: 'a::linorder. -atLeast k = lessThan k"
   114 apply (simp add: lessThan_def atLeast_def le_def, auto)
   115 done
   116 
   117 lemma atMost_iff [iff]: "(i: atMost k) = (i<=k)"
   118 by (simp add: atMost_def)
   119 
   120 lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}"
   121 by (blast intro: order_antisym)
   122 
   123 
   124 subsection {* Logical Equivalences for Set Inclusion and Equality *}
   125 
   126 lemma atLeast_subset_iff [iff]:
   127      "(atLeast x \<subseteq> atLeast y) = (y \<le> (x::'a::order))" 
   128 by (blast intro: order_trans) 
   129 
   130 lemma atLeast_eq_iff [iff]:
   131      "(atLeast x = atLeast y) = (x = (y::'a::linorder))" 
   132 by (blast intro: order_antisym order_trans)
   133 
   134 lemma greaterThan_subset_iff [iff]:
   135      "(greaterThan x \<subseteq> greaterThan y) = (y \<le> (x::'a::linorder))" 
   136 apply (auto simp add: greaterThan_def) 
   137  apply (subst linorder_not_less [symmetric], blast) 
   138 done
   139 
   140 lemma greaterThan_eq_iff [iff]:
   141      "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" 
   142 apply (rule iffI) 
   143  apply (erule equalityE) 
   144  apply (simp add: greaterThan_subset_iff order_antisym, simp) 
   145 done
   146 
   147 lemma atMost_subset_iff [iff]: "(atMost x \<subseteq> atMost y) = (x \<le> (y::'a::order))" 
   148 by (blast intro: order_trans)
   149 
   150 lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))" 
   151 by (blast intro: order_antisym order_trans)
   152 
   153 lemma lessThan_subset_iff [iff]:
   154      "(lessThan x \<subseteq> lessThan y) = (x \<le> (y::'a::linorder))" 
   155 apply (auto simp add: lessThan_def) 
   156  apply (subst linorder_not_less [symmetric], blast) 
   157 done
   158 
   159 lemma lessThan_eq_iff [iff]:
   160      "(lessThan x = lessThan y) = (x = (y::'a::linorder))" 
   161 apply (rule iffI) 
   162  apply (erule equalityE) 
   163  apply (simp add: lessThan_subset_iff order_antisym, simp) 
   164 done
   165 
   166 
   167 subsection {*Two-sided intervals*}
   168 
   169 text {* @{text greaterThanLessThan} *}
   170 
   171 lemma greaterThanLessThan_iff [simp]:
   172   "(i : {l<..<u}) = (l < i & i < u)"
   173 by (simp add: greaterThanLessThan_def)
   174 
   175 text {* @{text atLeastLessThan} *}
   176 
   177 lemma atLeastLessThan_iff [simp]:
   178   "(i : {l..<u}) = (l <= i & i < u)"
   179 by (simp add: atLeastLessThan_def)
   180 
   181 text {* @{text greaterThanAtMost} *}
   182 
   183 lemma greaterThanAtMost_iff [simp]:
   184   "(i : {l<..u}) = (l < i & i <= u)"
   185 by (simp add: greaterThanAtMost_def)
   186 
   187 text {* @{text atLeastAtMost} *}
   188 
   189 lemma atLeastAtMost_iff [simp]:
   190   "(i : {l..u}) = (l <= i & i <= u)"
   191 by (simp add: atLeastAtMost_def)
   192 
   193 text {* The above four lemmas could be declared as iffs.
   194   If we do so, a call to blast in Hyperreal/Star.ML, lemma @{text STAR_Int}
   195   seems to take forever (more than one hour). *}
   196 
   197 
   198 subsection {* Intervals of natural numbers *}
   199 
   200 subsubsection {* The Constant @{term lessThan} *}
   201 
   202 lemma lessThan_0 [simp]: "lessThan (0::nat) = {}"
   203 by (simp add: lessThan_def)
   204 
   205 lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)"
   206 by (simp add: lessThan_def less_Suc_eq, blast)
   207 
   208 lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k"
   209 by (simp add: lessThan_def atMost_def less_Suc_eq_le)
   210 
   211 lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV"
   212 by blast
   213 
   214 subsubsection {* The Constant @{term greaterThan} *}
   215 
   216 lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc"
   217 apply (simp add: greaterThan_def)
   218 apply (blast dest: gr0_conv_Suc [THEN iffD1])
   219 done
   220 
   221 lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}"
   222 apply (simp add: greaterThan_def)
   223 apply (auto elim: linorder_neqE)
   224 done
   225 
   226 lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}"
   227 by blast
   228 
   229 subsubsection {* The Constant @{term atLeast} *}
   230 
   231 lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV"
   232 by (unfold atLeast_def UNIV_def, simp)
   233 
   234 lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}"
   235 apply (simp add: atLeast_def)
   236 apply (simp add: Suc_le_eq)
   237 apply (simp add: order_le_less, blast)
   238 done
   239 
   240 lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k"
   241   by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le)
   242 
   243 lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV"
   244 by blast
   245 
   246 subsubsection {* The Constant @{term atMost} *}
   247 
   248 lemma atMost_0 [simp]: "atMost (0::nat) = {0}"
   249 by (simp add: atMost_def)
   250 
   251 lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)"
   252 apply (simp add: atMost_def)
   253 apply (simp add: less_Suc_eq order_le_less, blast)
   254 done
   255 
   256 lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV"
   257 by blast
   258 
   259 subsubsection {* The Constant @{term atLeastLessThan} *}
   260 
   261 text{*But not a simprule because some concepts are better left in terms
   262   of @{term atLeastLessThan}*}
   263 lemma atLeast0LessThan: "{0::nat..<n} = {..<n}"
   264 by(simp add:lessThan_def atLeastLessThan_def)
   265 
   266 lemma atLeastLessThan0 [simp]: "{m..<0::nat} = {}"
   267 by (simp add: atLeastLessThan_def)
   268 
   269 lemma atLeastLessThan_self [simp]: "{n::'a::order..<n} = {}"
   270 by (auto simp add: atLeastLessThan_def)
   271 
   272 lemma atLeastLessThan_empty: "n \<le> m ==> {m..<n::'a::order} = {}"
   273 by (auto simp add: atLeastLessThan_def)
   274 
   275 subsubsection {* Intervals of nats with @{term Suc} *}
   276 
   277 text{*Not a simprule because the RHS is too messy.*}
   278 lemma atLeastLessThanSuc:
   279     "{m..<Suc n} = (if m \<le> n then insert n {m..<n} else {})"
   280 by (auto simp add: atLeastLessThan_def) 
   281 
   282 lemma atLeastLessThan_singleton [simp]: "{m..<Suc m} = {m}" 
   283 by (auto simp add: atLeastLessThan_def)
   284 
   285 lemma atLeast_sum_LessThan [simp]: "{m + k..<k::nat} = {}"
   286 by (induct k, simp_all add: atLeastLessThanSuc)
   287 
   288 lemma atLeastSucLessThan [simp]: "{Suc n..<n} = {}"
   289 by (auto simp add: atLeastLessThan_def)
   290 
   291 lemma atLeastLessThanSuc_atLeastAtMost: "{l..<Suc u} = {l..u}"
   292   by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def)
   293 
   294 lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {l<..u}"  
   295   by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def 
   296     greaterThanAtMost_def)
   297 
   298 lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..<u} = {l<..<u}"  
   299   by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def 
   300     greaterThanLessThan_def)
   301 
   302 subsubsection {* Finiteness *}
   303 
   304 lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..<k}"
   305   by (induct k) (simp_all add: lessThan_Suc)
   306 
   307 lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}"
   308   by (induct k) (simp_all add: atMost_Suc)
   309 
   310 lemma finite_greaterThanLessThan [iff]:
   311   fixes l :: nat shows "finite {l<..<u}"
   312 by (simp add: greaterThanLessThan_def)
   313 
   314 lemma finite_atLeastLessThan [iff]:
   315   fixes l :: nat shows "finite {l..<u}"
   316 by (simp add: atLeastLessThan_def)
   317 
   318 lemma finite_greaterThanAtMost [iff]:
   319   fixes l :: nat shows "finite {l<..u}"
   320 by (simp add: greaterThanAtMost_def)
   321 
   322 lemma finite_atLeastAtMost [iff]:
   323   fixes l :: nat shows "finite {l..u}"
   324 by (simp add: atLeastAtMost_def)
   325 
   326 lemma bounded_nat_set_is_finite:
   327     "(ALL i:N. i < (n::nat)) ==> finite N"
   328   -- {* A bounded set of natural numbers is finite. *}
   329   apply (rule finite_subset)
   330    apply (rule_tac [2] finite_lessThan, auto)
   331   done
   332 
   333 subsubsection {* Cardinality *}
   334 
   335 lemma card_lessThan [simp]: "card {..<u} = u"
   336   by (induct_tac u, simp_all add: lessThan_Suc)
   337 
   338 lemma card_atMost [simp]: "card {..u} = Suc u"
   339   by (simp add: lessThan_Suc_atMost [THEN sym])
   340 
   341 lemma card_atLeastLessThan [simp]: "card {l..<u} = u - l"
   342   apply (subgoal_tac "card {l..<u} = card {..<u-l}")
   343   apply (erule ssubst, rule card_lessThan)
   344   apply (subgoal_tac "(%x. x + l) ` {..<u-l} = {l..<u}")
   345   apply (erule subst)
   346   apply (rule card_image)
   347   apply (rule finite_lessThan)
   348   apply (simp add: inj_on_def)
   349   apply (auto simp add: image_def atLeastLessThan_def lessThan_def)
   350   apply arith
   351   apply (rule_tac x = "x - l" in exI)
   352   apply arith
   353   done
   354 
   355 lemma card_atLeastAtMost [simp]: "card {l..u} = Suc u - l"  
   356   by (subst atLeastLessThanSuc_atLeastAtMost [THEN sym], simp)
   357 
   358 lemma card_greaterThanAtMost [simp]: "card {l<..u} = u - l" 
   359   by (subst atLeastSucAtMost_greaterThanAtMost [THEN sym], simp)
   360 
   361 lemma card_greaterThanLessThan [simp]: "card {l<..<u} = u - Suc l"
   362   by (subst atLeastSucLessThan_greaterThanLessThan [THEN sym], simp)
   363 
   364 subsection {* Intervals of integers *}
   365 
   366 lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..<u+1} = {l..(u::int)}"
   367   by (auto simp add: atLeastAtMost_def atLeastLessThan_def)
   368 
   369 lemma atLeastPlusOneAtMost_greaterThanAtMost_int: "{l+1..u} = {l<..(u::int)}"  
   370   by (auto simp add: atLeastAtMost_def greaterThanAtMost_def)
   371 
   372 lemma atLeastPlusOneLessThan_greaterThanLessThan_int: 
   373     "{l+1..<u} = {l<..<u::int}"  
   374   by (auto simp add: atLeastLessThan_def greaterThanLessThan_def)
   375 
   376 subsubsection {* Finiteness *}
   377 
   378 lemma image_atLeastZeroLessThan_int: "0 \<le> u ==> 
   379     {(0::int)..<u} = int ` {..<nat u}"
   380   apply (unfold image_def lessThan_def)
   381   apply auto
   382   apply (rule_tac x = "nat x" in exI)
   383   apply (auto simp add: zless_nat_conj zless_nat_eq_int_zless [THEN sym])
   384   done
   385 
   386 lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..<u}"
   387   apply (case_tac "0 \<le> u")
   388   apply (subst image_atLeastZeroLessThan_int, assumption)
   389   apply (rule finite_imageI)
   390   apply auto
   391   apply (subgoal_tac "{0..<u} = {}")
   392   apply auto
   393   done
   394 
   395 lemma image_atLeastLessThan_int_shift: 
   396     "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}"
   397   apply (auto simp add: image_def atLeastLessThan_iff)
   398   apply (rule_tac x = "x - l" in bexI)
   399   apply auto
   400   done
   401 
   402 lemma finite_atLeastLessThan_int [iff]: "finite {l..<u::int}"
   403   apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
   404   apply (erule subst)
   405   apply (rule finite_imageI)
   406   apply (rule finite_atLeastZeroLessThan_int)
   407   apply (rule image_atLeastLessThan_int_shift)
   408   done
   409 
   410 lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}" 
   411   by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp)
   412 
   413 lemma finite_greaterThanAtMost_int [iff]: "finite {l<..(u::int)}" 
   414   by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
   415 
   416 lemma finite_greaterThanLessThan_int [iff]: "finite {l<..<u::int}" 
   417   by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
   418 
   419 subsubsection {* Cardinality *}
   420 
   421 lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u"
   422   apply (case_tac "0 \<le> u")
   423   apply (subst image_atLeastZeroLessThan_int, assumption)
   424   apply (subst card_image)
   425   apply (auto simp add: inj_on_def)
   426   done
   427 
   428 lemma card_atLeastLessThan_int [simp]: "card {l..<u} = nat (u - l)"
   429   apply (subgoal_tac "card {l..<u} = card {0..<u-l}")
   430   apply (erule ssubst, rule card_atLeastZeroLessThan_int)
   431   apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
   432   apply (erule subst)
   433   apply (rule card_image)
   434   apply (rule finite_atLeastZeroLessThan_int)
   435   apply (simp add: inj_on_def)
   436   apply (rule image_atLeastLessThan_int_shift)
   437   done
   438 
   439 lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)"
   440   apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
   441   apply (auto simp add: compare_rls)
   442   done
   443 
   444 lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)" 
   445   by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
   446 
   447 lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))"
   448   by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
   449 
   450 
   451 subsection {*Lemmas useful with the summation operator setsum*}
   452 
   453 text {* For examples, see Algebra/poly/UnivPoly.thy *}
   454 
   455 subsubsection {* Disjoint Unions *}
   456 
   457 text {* Singletons and open intervals *}
   458 
   459 lemma ivl_disj_un_singleton:
   460   "{l::'a::linorder} Un {l<..} = {l..}"
   461   "{..<u} Un {u::'a::linorder} = {..u}"
   462   "(l::'a::linorder) < u ==> {l} Un {l<..<u} = {l..<u}"
   463   "(l::'a::linorder) < u ==> {l<..<u} Un {u} = {l<..u}"
   464   "(l::'a::linorder) <= u ==> {l} Un {l<..u} = {l..u}"
   465   "(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}"
   466 by auto
   467 
   468 text {* One- and two-sided intervals *}
   469 
   470 lemma ivl_disj_un_one:
   471   "(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}"
   472   "(l::'a::linorder) <= u ==> {..<l} Un {l..<u} = {..<u}"
   473   "(l::'a::linorder) <= u ==> {..l} Un {l<..u} = {..u}"
   474   "(l::'a::linorder) <= u ==> {..<l} Un {l..u} = {..u}"
   475   "(l::'a::linorder) <= u ==> {l<..u} Un {u<..} = {l<..}"
   476   "(l::'a::linorder) < u ==> {l<..<u} Un {u..} = {l<..}"
   477   "(l::'a::linorder) <= u ==> {l..u} Un {u<..} = {l..}"
   478   "(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}"
   479 by auto
   480 
   481 text {* Two- and two-sided intervals *}
   482 
   483 lemma ivl_disj_un_two:
   484   "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
   485   "[| (l::'a::linorder) <= m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}"
   486   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..<u} = {l..<u}"
   487   "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}"
   488   "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..u} = {l<..u}"
   489   "[| (l::'a::linorder) <= m; m <= u |] ==> {l<..m} Un {m<..u} = {l<..u}"
   490   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..u} = {l..u}"
   491   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m<..u} = {l..u}"
   492 by auto
   493 
   494 lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two
   495 
   496 subsubsection {* Disjoint Intersections *}
   497 
   498 text {* Singletons and open intervals *}
   499 
   500 lemma ivl_disj_int_singleton:
   501   "{l::'a::order} Int {l<..} = {}"
   502   "{..<u} Int {u} = {}"
   503   "{l} Int {l<..<u} = {}"
   504   "{l<..<u} Int {u} = {}"
   505   "{l} Int {l<..u} = {}"
   506   "{l..<u} Int {u} = {}"
   507   by simp+
   508 
   509 text {* One- and two-sided intervals *}
   510 
   511 lemma ivl_disj_int_one:
   512   "{..l::'a::order} Int {l<..<u} = {}"
   513   "{..<l} Int {l..<u} = {}"
   514   "{..l} Int {l<..u} = {}"
   515   "{..<l} Int {l..u} = {}"
   516   "{l<..u} Int {u<..} = {}"
   517   "{l<..<u} Int {u..} = {}"
   518   "{l..u} Int {u<..} = {}"
   519   "{l..<u} Int {u..} = {}"
   520   by auto
   521 
   522 text {* Two- and two-sided intervals *}
   523 
   524 lemma ivl_disj_int_two:
   525   "{l::'a::order<..<m} Int {m..<u} = {}"
   526   "{l<..m} Int {m<..<u} = {}"
   527   "{l..<m} Int {m..<u} = {}"
   528   "{l..m} Int {m<..<u} = {}"
   529   "{l<..<m} Int {m..u} = {}"
   530   "{l<..m} Int {m<..u} = {}"
   531   "{l..<m} Int {m..u} = {}"
   532   "{l..m} Int {m<..u} = {}"
   533   by auto
   534 
   535 lemmas ivl_disj_int = ivl_disj_int_singleton ivl_disj_int_one ivl_disj_int_two
   536 
   537 
   538 subsection {* Summation indexed over intervals *}
   539 
   540 syntax
   541   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
   542   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10)
   543   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<_./ _)" [0,0,10] 10)
   544 syntax (xsymbols)
   545   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
   546   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
   547   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
   548 syntax (HTML output)
   549   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
   550   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
   551   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
   552 syntax (latex output)
   553   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
   554  ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10)
   555   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
   556  ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10)
   557   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
   558  ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10)
   559 
   560 translations
   561   "\<Sum>x=a..b. t" == "setsum (%x. t) {a..b}"
   562   "\<Sum>x=a..<b. t" == "setsum (%x. t) {a..<b}"
   563   "\<Sum>i<n. t" == "setsum (\<lambda>i. t) {..<n}"
   564 
   565 text{* The above introduces some pretty alternative syntaxes for
   566 summation over intervals as shown on the left-hand side:
   567 \begin{center}
   568 \begin{tabular}{lll}
   569 Sets & Indexed & TeX\\
   570 @{term[source]"\<Sum>x\<in>{a..b}. e"} & @{term[source]"\<Sum>x=a..b. e"} & @{term"\<Sum>x=a..b. e"}\\
   571 @{term[source]"\<Sum>x\<in>{a..<b}. e"} & @{term[source]"\<Sum>x=a..<b. e"} & @{term"\<Sum>x=a..<b. e"}\\
   572 @{term[source]"\<Sum>x\<in>{..<b}. e"} & @{term[source]"\<Sum>x<b. e"} & @{term"\<Sum>x<b. e"}
   573 \end{tabular}
   574 \end{center}
   575 
   576 Note that for uniformity on @{typ nat} it is better to use
   577 @{term"\<Sum>x::nat=0..<n. e"} rather than @{text"\<Sum>x<n. e"}: @{text setsum} may
   578 not provide all lemmas available for @{term"{m..<n}"} also in the
   579 special form for @{term"{..<n}"}. *}
   580 
   581 
   582 lemma Summation_Suc[simp]: "(\<Sum>i < Suc n. b i) = b n + (\<Sum>i < n. b i)"
   583 by (simp add:lessThan_Suc)
   584 
   585 end