src/HOL/SetInterval.thy
author haftmann
Sat, 30 Jul 2011 08:24:46 +0200
changeset 44879 2e09299ce807
parent 44515 537ea3846f64
child 45761 22f665a2e91c
permissions -rw-r--r--
tuned proofs
     1 (*  Title:      HOL/SetInterval.thy
     2     Author:     Tobias Nipkow
     3     Author:     Clemens Ballarin
     4     Author:     Jeremy Avigad
     5 
     6 lessThan, greaterThan, atLeast, atMost and two-sided intervals
     7 *)
     8 
     9 header {* Set intervals *}
    10 
    11 theory SetInterval
    12 imports Int Nat_Transfer
    13 begin
    14 
    15 context ord
    16 begin
    17 
    18 definition
    19   lessThan    :: "'a => 'a set" ("(1{..<_})") where
    20   "{..<u} == {x. x < u}"
    21 
    22 definition
    23   atMost      :: "'a => 'a set" ("(1{.._})") where
    24   "{..u} == {x. x \<le> u}"
    25 
    26 definition
    27   greaterThan :: "'a => 'a set" ("(1{_<..})") where
    28   "{l<..} == {x. l<x}"
    29 
    30 definition
    31   atLeast     :: "'a => 'a set" ("(1{_..})") where
    32   "{l..} == {x. l\<le>x}"
    33 
    34 definition
    35   greaterThanLessThan :: "'a => 'a => 'a set"  ("(1{_<..<_})") where
    36   "{l<..<u} == {l<..} Int {..<u}"
    37 
    38 definition
    39   atLeastLessThan :: "'a => 'a => 'a set"      ("(1{_..<_})") where
    40   "{l..<u} == {l..} Int {..<u}"
    41 
    42 definition
    43   greaterThanAtMost :: "'a => 'a => 'a set"    ("(1{_<.._})") where
    44   "{l<..u} == {l<..} Int {..u}"
    45 
    46 definition
    47   atLeastAtMost :: "'a => 'a => 'a set"        ("(1{_.._})") where
    48   "{l..u} == {l..} Int {..u}"
    49 
    50 end
    51 
    52 
    53 text{* A note of warning when using @{term"{..<n}"} on type @{typ
    54 nat}: it is equivalent to @{term"{0::nat..<n}"} but some lemmas involving
    55 @{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *}
    56 
    57 syntax
    58   "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3UN _<=_./ _)" [0, 0, 10] 10)
    59   "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3UN _<_./ _)" [0, 0, 10] 10)
    60   "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3INT _<=_./ _)" [0, 0, 10] 10)
    61   "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3INT _<_./ _)" [0, 0, 10] 10)
    62 
    63 syntax (xsymbols)
    64   "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _\<le>_./ _)" [0, 0, 10] 10)
    65   "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _<_./ _)" [0, 0, 10] 10)
    66   "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _\<le>_./ _)" [0, 0, 10] 10)
    67   "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _<_./ _)" [0, 0, 10] 10)
    68 
    69 syntax (latex output)
    70   "_UNION_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ \<le> _)/ _)" [0, 0, 10] 10)
    71   "_UNION_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ < _)/ _)" [0, 0, 10] 10)
    72   "_INTER_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ \<le> _)/ _)" [0, 0, 10] 10)
    73   "_INTER_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ < _)/ _)" [0, 0, 10] 10)
    74 
    75 translations
    76   "UN i<=n. A"  == "UN i:{..n}. A"
    77   "UN i<n. A"   == "UN i:{..<n}. A"
    78   "INT i<=n. A" == "INT i:{..n}. A"
    79   "INT i<n. A"  == "INT i:{..<n}. A"
    80 
    81 
    82 subsection {* Various equivalences *}
    83 
    84 lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i<k)"
    85 by (simp add: lessThan_def)
    86 
    87 lemma Compl_lessThan [simp]:
    88     "!!k:: 'a::linorder. -lessThan k = atLeast k"
    89 apply (auto simp add: lessThan_def atLeast_def)
    90 done
    91 
    92 lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}"
    93 by auto
    94 
    95 lemma (in ord) greaterThan_iff [iff]: "(i: greaterThan k) = (k<i)"
    96 by (simp add: greaterThan_def)
    97 
    98 lemma Compl_greaterThan [simp]:
    99     "!!k:: 'a::linorder. -greaterThan k = atMost k"
   100   by (auto simp add: greaterThan_def atMost_def)
   101 
   102 lemma Compl_atMost [simp]: "!!k:: 'a::linorder. -atMost k = greaterThan k"
   103 apply (subst Compl_greaterThan [symmetric])
   104 apply (rule double_complement)
   105 done
   106 
   107 lemma (in ord) atLeast_iff [iff]: "(i: atLeast k) = (k<=i)"
   108 by (simp add: atLeast_def)
   109 
   110 lemma Compl_atLeast [simp]:
   111     "!!k:: 'a::linorder. -atLeast k = lessThan k"
   112   by (auto simp add: lessThan_def atLeast_def)
   113 
   114 lemma (in ord) atMost_iff [iff]: "(i: atMost k) = (i<=k)"
   115 by (simp add: atMost_def)
   116 
   117 lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}"
   118 by (blast intro: order_antisym)
   119 
   120 
   121 subsection {* Logical Equivalences for Set Inclusion and Equality *}
   122 
   123 lemma atLeast_subset_iff [iff]:
   124      "(atLeast x \<subseteq> atLeast y) = (y \<le> (x::'a::order))"
   125 by (blast intro: order_trans)
   126 
   127 lemma atLeast_eq_iff [iff]:
   128      "(atLeast x = atLeast y) = (x = (y::'a::linorder))"
   129 by (blast intro: order_antisym order_trans)
   130 
   131 lemma greaterThan_subset_iff [iff]:
   132      "(greaterThan x \<subseteq> greaterThan y) = (y \<le> (x::'a::linorder))"
   133 apply (auto simp add: greaterThan_def)
   134  apply (subst linorder_not_less [symmetric], blast)
   135 done
   136 
   137 lemma greaterThan_eq_iff [iff]:
   138      "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))"
   139 apply (rule iffI)
   140  apply (erule equalityE)
   141  apply simp_all
   142 done
   143 
   144 lemma atMost_subset_iff [iff]: "(atMost x \<subseteq> atMost y) = (x \<le> (y::'a::order))"
   145 by (blast intro: order_trans)
   146 
   147 lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))"
   148 by (blast intro: order_antisym order_trans)
   149 
   150 lemma lessThan_subset_iff [iff]:
   151      "(lessThan x \<subseteq> lessThan y) = (x \<le> (y::'a::linorder))"
   152 apply (auto simp add: lessThan_def)
   153  apply (subst linorder_not_less [symmetric], blast)
   154 done
   155 
   156 lemma lessThan_eq_iff [iff]:
   157      "(lessThan x = lessThan y) = (x = (y::'a::linorder))"
   158 apply (rule iffI)
   159  apply (erule equalityE)
   160  apply simp_all
   161 done
   162 
   163 lemma lessThan_strict_subset_iff:
   164   fixes m n :: "'a::linorder"
   165   shows "{..<m} < {..<n} \<longleftrightarrow> m < n"
   166   by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq)
   167 
   168 subsection {*Two-sided intervals*}
   169 
   170 context ord
   171 begin
   172 
   173 lemma greaterThanLessThan_iff [simp,no_atp]:
   174   "(i : {l<..<u}) = (l < i & i < u)"
   175 by (simp add: greaterThanLessThan_def)
   176 
   177 lemma atLeastLessThan_iff [simp,no_atp]:
   178   "(i : {l..<u}) = (l <= i & i < u)"
   179 by (simp add: atLeastLessThan_def)
   180 
   181 lemma greaterThanAtMost_iff [simp,no_atp]:
   182   "(i : {l<..u}) = (l < i & i <= u)"
   183 by (simp add: greaterThanAtMost_def)
   184 
   185 lemma atLeastAtMost_iff [simp,no_atp]:
   186   "(i : {l..u}) = (l <= i & i <= u)"
   187 by (simp add: atLeastAtMost_def)
   188 
   189 text {* The above four lemmas could be declared as iffs. Unfortunately this
   190 breaks many proofs. Since it only helps blast, it is better to leave well
   191 alone *}
   192 
   193 end
   194 
   195 subsubsection{* Emptyness, singletons, subset *}
   196 
   197 context order
   198 begin
   199 
   200 lemma atLeastatMost_empty[simp]:
   201   "b < a \<Longrightarrow> {a..b} = {}"
   202 by(auto simp: atLeastAtMost_def atLeast_def atMost_def)
   203 
   204 lemma atLeastatMost_empty_iff[simp]:
   205   "{a..b} = {} \<longleftrightarrow> (~ a <= b)"
   206 by auto (blast intro: order_trans)
   207 
   208 lemma atLeastatMost_empty_iff2[simp]:
   209   "{} = {a..b} \<longleftrightarrow> (~ a <= b)"
   210 by auto (blast intro: order_trans)
   211 
   212 lemma atLeastLessThan_empty[simp]:
   213   "b <= a \<Longrightarrow> {a..<b} = {}"
   214 by(auto simp: atLeastLessThan_def)
   215 
   216 lemma atLeastLessThan_empty_iff[simp]:
   217   "{a..<b} = {} \<longleftrightarrow> (~ a < b)"
   218 by auto (blast intro: le_less_trans)
   219 
   220 lemma atLeastLessThan_empty_iff2[simp]:
   221   "{} = {a..<b} \<longleftrightarrow> (~ a < b)"
   222 by auto (blast intro: le_less_trans)
   223 
   224 lemma greaterThanAtMost_empty[simp]: "l \<le> k ==> {k<..l} = {}"
   225 by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def)
   226 
   227 lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \<longleftrightarrow> ~ k < l"
   228 by auto (blast intro: less_le_trans)
   229 
   230 lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \<longleftrightarrow> ~ k < l"
   231 by auto (blast intro: less_le_trans)
   232 
   233 lemma greaterThanLessThan_empty[simp]:"l \<le> k ==> {k<..<l} = {}"
   234 by(auto simp:greaterThanLessThan_def greaterThan_def lessThan_def)
   235 
   236 lemma atLeastAtMost_singleton [simp]: "{a..a} = {a}"
   237 by (auto simp add: atLeastAtMost_def atMost_def atLeast_def)
   238 
   239 lemma atLeastAtMost_singleton': "a = b \<Longrightarrow> {a .. b} = {a}" by simp
   240 
   241 lemma atLeastatMost_subset_iff[simp]:
   242   "{a..b} <= {c..d} \<longleftrightarrow> (~ a <= b) | c <= a & b <= d"
   243 unfolding atLeastAtMost_def atLeast_def atMost_def
   244 by (blast intro: order_trans)
   245 
   246 lemma atLeastatMost_psubset_iff:
   247   "{a..b} < {c..d} \<longleftrightarrow>
   248    ((~ a <= b) | c <= a & b <= d & (c < a | b < d))  &  c <= d"
   249 by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans)
   250 
   251 lemma atLeastAtMost_singleton_iff[simp]:
   252   "{a .. b} = {c} \<longleftrightarrow> a = b \<and> b = c"
   253 proof
   254   assume "{a..b} = {c}"
   255   hence "\<not> (\<not> a \<le> b)" unfolding atLeastatMost_empty_iff[symmetric] by simp
   256   moreover with `{a..b} = {c}` have "c \<le> a \<and> b \<le> c" by auto
   257   ultimately show "a = b \<and> b = c" by auto
   258 qed simp
   259 
   260 end
   261 
   262 context dense_linorder
   263 begin
   264 
   265 lemma greaterThanLessThan_empty_iff[simp]:
   266   "{ a <..< b } = {} \<longleftrightarrow> b \<le> a"
   267   using dense[of a b] by (cases "a < b") auto
   268 
   269 lemma greaterThanLessThan_empty_iff2[simp]:
   270   "{} = { a <..< b } \<longleftrightarrow> b \<le> a"
   271   using dense[of a b] by (cases "a < b") auto
   272 
   273 lemma atLeastLessThan_subseteq_atLeastAtMost_iff:
   274   "{a ..< b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
   275   using dense[of "max a d" "b"]
   276   by (force simp: subset_eq Ball_def not_less[symmetric])
   277 
   278 lemma greaterThanAtMost_subseteq_atLeastAtMost_iff:
   279   "{a <.. b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
   280   using dense[of "a" "min c b"]
   281   by (force simp: subset_eq Ball_def not_less[symmetric])
   282 
   283 lemma greaterThanLessThan_subseteq_atLeastAtMost_iff:
   284   "{a <..< b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
   285   using dense[of "a" "min c b"] dense[of "max a d" "b"]
   286   by (force simp: subset_eq Ball_def not_less[symmetric])
   287 
   288 lemma atLeastAtMost_subseteq_atLeastLessThan_iff:
   289   "{a .. b} \<subseteq> { c ..< d } \<longleftrightarrow> (a \<le> b \<longrightarrow> c \<le> a \<and> b < d)"
   290   using dense[of "max a d" "b"]
   291   by (force simp: subset_eq Ball_def not_less[symmetric])
   292 
   293 lemma greaterThanAtMost_subseteq_atLeastLessThan_iff:
   294   "{a <.. b} \<subseteq> { c ..< d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b < d)"
   295   using dense[of "a" "min c b"]
   296   by (force simp: subset_eq Ball_def not_less[symmetric])
   297 
   298 lemma greaterThanLessThan_subseteq_atLeastLessThan_iff:
   299   "{a <..< b} \<subseteq> { c ..< d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
   300   using dense[of "a" "min c b"] dense[of "max a d" "b"]
   301   by (force simp: subset_eq Ball_def not_less[symmetric])
   302 
   303 end
   304 
   305 lemma (in linorder) atLeastLessThan_subset_iff:
   306   "{a..<b} <= {c..<d} \<Longrightarrow> b <= a | c<=a & b<=d"
   307 apply (auto simp:subset_eq Ball_def)
   308 apply(frule_tac x=a in spec)
   309 apply(erule_tac x=d in allE)
   310 apply (simp add: less_imp_le)
   311 done
   312 
   313 lemma atLeastLessThan_inj:
   314   fixes a b c d :: "'a::linorder"
   315   assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d"
   316   shows "a = c" "b = d"
   317 using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le linorder_antisym_conv2 subset_refl)+
   318 
   319 lemma atLeastLessThan_eq_iff:
   320   fixes a b c d :: "'a::linorder"
   321   assumes "a < b" "c < d"
   322   shows "{a ..< b} = {c ..< d} \<longleftrightarrow> a = c \<and> b = d"
   323   using atLeastLessThan_inj assms by auto
   324 
   325 subsubsection {* Intersection *}
   326 
   327 context linorder
   328 begin
   329 
   330 lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}"
   331 by auto
   332 
   333 lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}"
   334 by auto
   335 
   336 lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}"
   337 by auto
   338 
   339 lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}"
   340 by auto
   341 
   342 lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}"
   343 by auto
   344 
   345 lemma Int_atLeastLessThan[simp]: "{a..<b} Int {c..<d} = {max a c ..< min b d}"
   346 by auto
   347 
   348 lemma Int_greaterThanAtMost[simp]: "{a<..b} Int {c<..d} = {max a c <.. min b d}"
   349 by auto
   350 
   351 lemma Int_greaterThanLessThan[simp]: "{a<..<b} Int {c<..<d} = {max a c <..< min b d}"
   352 by auto
   353 
   354 end
   355 
   356 
   357 subsection {* Intervals of natural numbers *}
   358 
   359 subsubsection {* The Constant @{term lessThan} *}
   360 
   361 lemma lessThan_0 [simp]: "lessThan (0::nat) = {}"
   362 by (simp add: lessThan_def)
   363 
   364 lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)"
   365 by (simp add: lessThan_def less_Suc_eq, blast)
   366 
   367 text {* The following proof is convenient in induction proofs where
   368 new elements get indices at the beginning. So it is used to transform
   369 @{term "{..<Suc n}"} to @{term "0::nat"} and @{term "{..< n}"}. *}
   370 
   371 lemma lessThan_Suc_eq_insert_0: "{..<Suc n} = insert 0 (Suc ` {..<n})"
   372 proof safe
   373   fix x assume "x < Suc n" "x \<notin> Suc ` {..<n}"
   374   then have "x \<noteq> Suc (x - 1)" by auto
   375   with `x < Suc n` show "x = 0" by auto
   376 qed
   377 
   378 lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k"
   379 by (simp add: lessThan_def atMost_def less_Suc_eq_le)
   380 
   381 lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV"
   382 by blast
   383 
   384 subsubsection {* The Constant @{term greaterThan} *}
   385 
   386 lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc"
   387 apply (simp add: greaterThan_def)
   388 apply (blast dest: gr0_conv_Suc [THEN iffD1])
   389 done
   390 
   391 lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}"
   392 apply (simp add: greaterThan_def)
   393 apply (auto elim: linorder_neqE)
   394 done
   395 
   396 lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}"
   397 by blast
   398 
   399 subsubsection {* The Constant @{term atLeast} *}
   400 
   401 lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV"
   402 by (unfold atLeast_def UNIV_def, simp)
   403 
   404 lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}"
   405 apply (simp add: atLeast_def)
   406 apply (simp add: Suc_le_eq)
   407 apply (simp add: order_le_less, blast)
   408 done
   409 
   410 lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k"
   411   by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le)
   412 
   413 lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV"
   414 by blast
   415 
   416 subsubsection {* The Constant @{term atMost} *}
   417 
   418 lemma atMost_0 [simp]: "atMost (0::nat) = {0}"
   419 by (simp add: atMost_def)
   420 
   421 lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)"
   422 apply (simp add: atMost_def)
   423 apply (simp add: less_Suc_eq order_le_less, blast)
   424 done
   425 
   426 lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV"
   427 by blast
   428 
   429 subsubsection {* The Constant @{term atLeastLessThan} *}
   430 
   431 text{*The orientation of the following 2 rules is tricky. The lhs is
   432 defined in terms of the rhs.  Hence the chosen orientation makes sense
   433 in this theory --- the reverse orientation complicates proofs (eg
   434 nontermination). But outside, when the definition of the lhs is rarely
   435 used, the opposite orientation seems preferable because it reduces a
   436 specific concept to a more general one. *}
   437 
   438 lemma atLeast0LessThan: "{0::nat..<n} = {..<n}"
   439 by(simp add:lessThan_def atLeastLessThan_def)
   440 
   441 lemma atLeast0AtMost: "{0..n::nat} = {..n}"
   442 by(simp add:atMost_def atLeastAtMost_def)
   443 
   444 declare atLeast0LessThan[symmetric, code_unfold]
   445         atLeast0AtMost[symmetric, code_unfold]
   446 
   447 lemma atLeastLessThan0: "{m..<0::nat} = {}"
   448 by (simp add: atLeastLessThan_def)
   449 
   450 subsubsection {* Intervals of nats with @{term Suc} *}
   451 
   452 text{*Not a simprule because the RHS is too messy.*}
   453 lemma atLeastLessThanSuc:
   454     "{m..<Suc n} = (if m \<le> n then insert n {m..<n} else {})"
   455 by (auto simp add: atLeastLessThan_def)
   456 
   457 lemma atLeastLessThan_singleton [simp]: "{m..<Suc m} = {m}"
   458 by (auto simp add: atLeastLessThan_def)
   459 (*
   460 lemma atLeast_sum_LessThan [simp]: "{m + k..<k::nat} = {}"
   461 by (induct k, simp_all add: atLeastLessThanSuc)
   462 
   463 lemma atLeastSucLessThan [simp]: "{Suc n..<n} = {}"
   464 by (auto simp add: atLeastLessThan_def)
   465 *)
   466 lemma atLeastLessThanSuc_atLeastAtMost: "{l..<Suc u} = {l..u}"
   467   by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def)
   468 
   469 lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {l<..u}"
   470   by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def
   471     greaterThanAtMost_def)
   472 
   473 lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..<u} = {l<..<u}"
   474   by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def
   475     greaterThanLessThan_def)
   476 
   477 lemma atLeastAtMostSuc_conv: "m \<le> Suc n \<Longrightarrow> {m..Suc n} = insert (Suc n) {m..n}"
   478 by (auto simp add: atLeastAtMost_def)
   479 
   480 text {* The analogous result is useful on @{typ int}: *}
   481 (* here, because we don't have an own int section *)
   482 lemma atLeastAtMostPlus1_int_conv:
   483   "m <= 1+n \<Longrightarrow> {m..1+n} = insert (1+n) {m..n::int}"
   484   by (auto intro: set_eqI)
   485 
   486 lemma atLeastLessThan_add_Un: "i \<le> j \<Longrightarrow> {i..<j+k} = {i..<j} \<union> {j..<j+k::nat}"
   487   apply (induct k) 
   488   apply (simp_all add: atLeastLessThanSuc)   
   489   done
   490 
   491 subsubsection {* Image *}
   492 
   493 lemma image_add_atLeastAtMost:
   494   "(%n::nat. n+k) ` {i..j} = {i+k..j+k}" (is "?A = ?B")
   495 proof
   496   show "?A \<subseteq> ?B" by auto
   497 next
   498   show "?B \<subseteq> ?A"
   499   proof
   500     fix n assume a: "n : ?B"
   501     hence "n - k : {i..j}" by auto
   502     moreover have "n = (n - k) + k" using a by auto
   503     ultimately show "n : ?A" by blast
   504   qed
   505 qed
   506 
   507 lemma image_add_atLeastLessThan:
   508   "(%n::nat. n+k) ` {i..<j} = {i+k..<j+k}" (is "?A = ?B")
   509 proof
   510   show "?A \<subseteq> ?B" by auto
   511 next
   512   show "?B \<subseteq> ?A"
   513   proof
   514     fix n assume a: "n : ?B"
   515     hence "n - k : {i..<j}" by auto
   516     moreover have "n = (n - k) + k" using a by auto
   517     ultimately show "n : ?A" by blast
   518   qed
   519 qed
   520 
   521 corollary image_Suc_atLeastAtMost[simp]:
   522   "Suc ` {i..j} = {Suc i..Suc j}"
   523 using image_add_atLeastAtMost[where k="Suc 0"] by simp
   524 
   525 corollary image_Suc_atLeastLessThan[simp]:
   526   "Suc ` {i..<j} = {Suc i..<Suc j}"
   527 using image_add_atLeastLessThan[where k="Suc 0"] by simp
   528 
   529 lemma image_add_int_atLeastLessThan:
   530     "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}"
   531   apply (auto simp add: image_def)
   532   apply (rule_tac x = "x - l" in bexI)
   533   apply auto
   534   done
   535 
   536 lemma image_minus_const_atLeastLessThan_nat:
   537   fixes c :: nat
   538   shows "(\<lambda>i. i - c) ` {x ..< y} =
   539       (if c < y then {x - c ..< y - c} else if x < y then {0} else {})"
   540     (is "_ = ?right")
   541 proof safe
   542   fix a assume a: "a \<in> ?right"
   543   show "a \<in> (\<lambda>i. i - c) ` {x ..< y}"
   544   proof cases
   545     assume "c < y" with a show ?thesis
   546       by (auto intro!: image_eqI[of _ _ "a + c"])
   547   next
   548     assume "\<not> c < y" with a show ?thesis
   549       by (auto intro!: image_eqI[of _ _ x] split: split_if_asm)
   550   qed
   551 qed auto
   552 
   553 context ordered_ab_group_add
   554 begin
   555 
   556 lemma
   557   fixes x :: 'a
   558   shows image_uminus_greaterThan[simp]: "uminus ` {x<..} = {..<-x}"
   559   and image_uminus_atLeast[simp]: "uminus ` {x..} = {..-x}"
   560 proof safe
   561   fix y assume "y < -x"
   562   hence *:  "x < -y" using neg_less_iff_less[of "-y" x] by simp
   563   have "- (-y) \<in> uminus ` {x<..}"
   564     by (rule imageI) (simp add: *)
   565   thus "y \<in> uminus ` {x<..}" by simp
   566 next
   567   fix y assume "y \<le> -x"
   568   have "- (-y) \<in> uminus ` {x..}"
   569     by (rule imageI) (insert `y \<le> -x`[THEN le_imp_neg_le], simp)
   570   thus "y \<in> uminus ` {x..}" by simp
   571 qed simp_all
   572 
   573 lemma
   574   fixes x :: 'a
   575   shows image_uminus_lessThan[simp]: "uminus ` {..<x} = {-x<..}"
   576   and image_uminus_atMost[simp]: "uminus ` {..x} = {-x..}"
   577 proof -
   578   have "uminus ` {..<x} = uminus ` uminus ` {-x<..}"
   579     and "uminus ` {..x} = uminus ` uminus ` {-x..}" by simp_all
   580   thus "uminus ` {..<x} = {-x<..}" and "uminus ` {..x} = {-x..}"
   581     by (simp_all add: image_image
   582         del: image_uminus_greaterThan image_uminus_atLeast)
   583 qed
   584 
   585 lemma
   586   fixes x :: 'a
   587   shows image_uminus_atLeastAtMost[simp]: "uminus ` {x..y} = {-y..-x}"
   588   and image_uminus_greaterThanAtMost[simp]: "uminus ` {x<..y} = {-y..<-x}"
   589   and image_uminus_atLeastLessThan[simp]: "uminus ` {x..<y} = {-y<..-x}"
   590   and image_uminus_greaterThanLessThan[simp]: "uminus ` {x<..<y} = {-y<..<-x}"
   591   by (simp_all add: atLeastAtMost_def greaterThanAtMost_def atLeastLessThan_def
   592       greaterThanLessThan_def image_Int[OF inj_uminus] Int_commute)
   593 end
   594 
   595 subsubsection {* Finiteness *}
   596 
   597 lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..<k}"
   598   by (induct k) (simp_all add: lessThan_Suc)
   599 
   600 lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}"
   601   by (induct k) (simp_all add: atMost_Suc)
   602 
   603 lemma finite_greaterThanLessThan [iff]:
   604   fixes l :: nat shows "finite {l<..<u}"
   605 by (simp add: greaterThanLessThan_def)
   606 
   607 lemma finite_atLeastLessThan [iff]:
   608   fixes l :: nat shows "finite {l..<u}"
   609 by (simp add: atLeastLessThan_def)
   610 
   611 lemma finite_greaterThanAtMost [iff]:
   612   fixes l :: nat shows "finite {l<..u}"
   613 by (simp add: greaterThanAtMost_def)
   614 
   615 lemma finite_atLeastAtMost [iff]:
   616   fixes l :: nat shows "finite {l..u}"
   617 by (simp add: atLeastAtMost_def)
   618 
   619 text {* A bounded set of natural numbers is finite. *}
   620 lemma bounded_nat_set_is_finite:
   621   "(ALL i:N. i < (n::nat)) ==> finite N"
   622 apply (rule finite_subset)
   623  apply (rule_tac [2] finite_lessThan, auto)
   624 done
   625 
   626 text {* A set of natural numbers is finite iff it is bounded. *}
   627 lemma finite_nat_set_iff_bounded:
   628   "finite(N::nat set) = (EX m. ALL n:N. n<m)" (is "?F = ?B")
   629 proof
   630   assume f:?F  show ?B
   631     using Max_ge[OF `?F`, simplified less_Suc_eq_le[symmetric]] by blast
   632 next
   633   assume ?B show ?F using `?B` by(blast intro:bounded_nat_set_is_finite)
   634 qed
   635 
   636 lemma finite_nat_set_iff_bounded_le:
   637   "finite(N::nat set) = (EX m. ALL n:N. n<=m)"
   638 apply(simp add:finite_nat_set_iff_bounded)
   639 apply(blast dest:less_imp_le_nat le_imp_less_Suc)
   640 done
   641 
   642 lemma finite_less_ub:
   643      "!!f::nat=>nat. (!!n. n \<le> f n) ==> finite {n. f n \<le> u}"
   644 by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans)
   645 
   646 text{* Any subset of an interval of natural numbers the size of the
   647 subset is exactly that interval. *}
   648 
   649 lemma subset_card_intvl_is_intvl:
   650   "A <= {k..<k+card A} \<Longrightarrow> A = {k..<k+card A}" (is "PROP ?P")
   651 proof cases
   652   assume "finite A"
   653   thus "PROP ?P"
   654   proof(induct A rule:finite_linorder_max_induct)
   655     case empty thus ?case by auto
   656   next
   657     case (insert b A)
   658     moreover hence "b ~: A" by auto
   659     moreover have "A <= {k..<k+card A}" and "b = k+card A"
   660       using `b ~: A` insert by fastsimp+
   661     ultimately show ?case by auto
   662   qed
   663 next
   664   assume "~finite A" thus "PROP ?P" by simp
   665 qed
   666 
   667 
   668 subsubsection {* Proving Inclusions and Equalities between Unions *}
   669 
   670 lemma UN_le_eq_Un0:
   671   "(\<Union>i\<le>n::nat. M i) = (\<Union>i\<in>{1..n}. M i) \<union> M 0" (is "?A = ?B")
   672 proof
   673   show "?A <= ?B"
   674   proof
   675     fix x assume "x : ?A"
   676     then obtain i where i: "i\<le>n" "x : M i" by auto
   677     show "x : ?B"
   678     proof(cases i)
   679       case 0 with i show ?thesis by simp
   680     next
   681       case (Suc j) with i show ?thesis by auto
   682     qed
   683   qed
   684 next
   685   show "?B <= ?A" by auto
   686 qed
   687 
   688 lemma UN_le_add_shift:
   689   "(\<Union>i\<le>n::nat. M(i+k)) = (\<Union>i\<in>{k..n+k}. M i)" (is "?A = ?B")
   690 proof
   691   show "?A <= ?B" by fastsimp
   692 next
   693   show "?B <= ?A"
   694   proof
   695     fix x assume "x : ?B"
   696     then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto
   697     hence "i-k\<le>n & x : M((i-k)+k)" by auto
   698     thus "x : ?A" by blast
   699   qed
   700 qed
   701 
   702 lemma UN_UN_finite_eq: "(\<Union>n::nat. \<Union>i\<in>{0..<n}. A i) = (\<Union>n. A n)"
   703   by (auto simp add: atLeast0LessThan) 
   704 
   705 lemma UN_finite_subset: "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> C) \<Longrightarrow> (\<Union>n. A n) \<subseteq> C"
   706   by (subst UN_UN_finite_eq [symmetric]) blast
   707 
   708 lemma UN_finite2_subset: 
   709      "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> (\<Union>i\<in>{0..<n+k}. B i)) \<Longrightarrow> (\<Union>n. A n) \<subseteq> (\<Union>n. B n)"
   710   apply (rule UN_finite_subset)
   711   apply (subst UN_UN_finite_eq [symmetric, of B]) 
   712   apply blast
   713   done
   714 
   715 lemma UN_finite2_eq:
   716   "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) = (\<Union>i\<in>{0..<n+k}. B i)) \<Longrightarrow> (\<Union>n. A n) = (\<Union>n. B n)"
   717   apply (rule subset_antisym)
   718    apply (rule UN_finite2_subset, blast)
   719  apply (rule UN_finite2_subset [where k=k])
   720  apply (force simp add: atLeastLessThan_add_Un [of 0])
   721  done
   722 
   723 
   724 subsubsection {* Cardinality *}
   725 
   726 lemma card_lessThan [simp]: "card {..<u} = u"
   727   by (induct u, simp_all add: lessThan_Suc)
   728 
   729 lemma card_atMost [simp]: "card {..u} = Suc u"
   730   by (simp add: lessThan_Suc_atMost [THEN sym])
   731 
   732 lemma card_atLeastLessThan [simp]: "card {l..<u} = u - l"
   733   apply (subgoal_tac "card {l..<u} = card {..<u-l}")
   734   apply (erule ssubst, rule card_lessThan)
   735   apply (subgoal_tac "(%x. x + l) ` {..<u-l} = {l..<u}")
   736   apply (erule subst)
   737   apply (rule card_image)
   738   apply (simp add: inj_on_def)
   739   apply (auto simp add: image_def atLeastLessThan_def lessThan_def)
   740   apply (rule_tac x = "x - l" in exI)
   741   apply arith
   742   done
   743 
   744 lemma card_atLeastAtMost [simp]: "card {l..u} = Suc u - l"
   745   by (subst atLeastLessThanSuc_atLeastAtMost [THEN sym], simp)
   746 
   747 lemma card_greaterThanAtMost [simp]: "card {l<..u} = u - l"
   748   by (subst atLeastSucAtMost_greaterThanAtMost [THEN sym], simp)
   749 
   750 lemma card_greaterThanLessThan [simp]: "card {l<..<u} = u - Suc l"
   751   by (subst atLeastSucLessThan_greaterThanLessThan [THEN sym], simp)
   752 
   753 lemma ex_bij_betw_nat_finite:
   754   "finite M \<Longrightarrow> \<exists>h. bij_betw h {0..<card M} M"
   755 apply(drule finite_imp_nat_seg_image_inj_on)
   756 apply(auto simp:atLeast0LessThan[symmetric] lessThan_def[symmetric] card_image bij_betw_def)
   757 done
   758 
   759 lemma ex_bij_betw_finite_nat:
   760   "finite M \<Longrightarrow> \<exists>h. bij_betw h M {0..<card M}"
   761 by (blast dest: ex_bij_betw_nat_finite bij_betw_inv)
   762 
   763 lemma finite_same_card_bij:
   764   "finite A \<Longrightarrow> finite B \<Longrightarrow> card A = card B \<Longrightarrow> EX h. bij_betw h A B"
   765 apply(drule ex_bij_betw_finite_nat)
   766 apply(drule ex_bij_betw_nat_finite)
   767 apply(auto intro!:bij_betw_trans)
   768 done
   769 
   770 lemma ex_bij_betw_nat_finite_1:
   771   "finite M \<Longrightarrow> \<exists>h. bij_betw h {1 .. card M} M"
   772 by (rule finite_same_card_bij) auto
   773 
   774 lemma bij_betw_iff_card:
   775   assumes FIN: "finite A" and FIN': "finite B"
   776   shows BIJ: "(\<exists>f. bij_betw f A B) \<longleftrightarrow> (card A = card B)"
   777 using assms
   778 proof(auto simp add: bij_betw_same_card)
   779   assume *: "card A = card B"
   780   obtain f where "bij_betw f A {0 ..< card A}"
   781   using FIN ex_bij_betw_finite_nat by blast
   782   moreover obtain g where "bij_betw g {0 ..< card B} B"
   783   using FIN' ex_bij_betw_nat_finite by blast
   784   ultimately have "bij_betw (g o f) A B"
   785   using * by (auto simp add: bij_betw_trans)
   786   thus "(\<exists>f. bij_betw f A B)" by blast
   787 qed
   788 
   789 lemma inj_on_iff_card_le:
   790   assumes FIN: "finite A" and FIN': "finite B"
   791   shows "(\<exists>f. inj_on f A \<and> f ` A \<le> B) = (card A \<le> card B)"
   792 proof (safe intro!: card_inj_on_le)
   793   assume *: "card A \<le> card B"
   794   obtain f where 1: "inj_on f A" and 2: "f ` A = {0 ..< card A}"
   795   using FIN ex_bij_betw_finite_nat unfolding bij_betw_def by force
   796   moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B"
   797   using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force
   798   ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force
   799   hence "inj_on (g o f) A" using 1 comp_inj_on by blast
   800   moreover
   801   {have "{0 ..< card A} \<le> {0 ..< card B}" using * by force
   802    with 2 have "f ` A  \<le> {0 ..< card B}" by blast
   803    hence "(g o f) ` A \<le> B" unfolding comp_def using 3 by force
   804   }
   805   ultimately show "(\<exists>f. inj_on f A \<and> f ` A \<le> B)" by blast
   806 qed (insert assms, auto)
   807 
   808 subsection {* Intervals of integers *}
   809 
   810 lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..<u+1} = {l..(u::int)}"
   811   by (auto simp add: atLeastAtMost_def atLeastLessThan_def)
   812 
   813 lemma atLeastPlusOneAtMost_greaterThanAtMost_int: "{l+1..u} = {l<..(u::int)}"
   814   by (auto simp add: atLeastAtMost_def greaterThanAtMost_def)
   815 
   816 lemma atLeastPlusOneLessThan_greaterThanLessThan_int:
   817     "{l+1..<u} = {l<..<u::int}"
   818   by (auto simp add: atLeastLessThan_def greaterThanLessThan_def)
   819 
   820 subsubsection {* Finiteness *}
   821 
   822 lemma image_atLeastZeroLessThan_int: "0 \<le> u ==>
   823     {(0::int)..<u} = int ` {..<nat u}"
   824   apply (unfold image_def lessThan_def)
   825   apply auto
   826   apply (rule_tac x = "nat x" in exI)
   827   apply (auto simp add: zless_nat_eq_int_zless [THEN sym])
   828   done
   829 
   830 lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..<u}"
   831   apply (case_tac "0 \<le> u")
   832   apply (subst image_atLeastZeroLessThan_int, assumption)
   833   apply (rule finite_imageI)
   834   apply auto
   835   done
   836 
   837 lemma finite_atLeastLessThan_int [iff]: "finite {l..<u::int}"
   838   apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
   839   apply (erule subst)
   840   apply (rule finite_imageI)
   841   apply (rule finite_atLeastZeroLessThan_int)
   842   apply (rule image_add_int_atLeastLessThan)
   843   done
   844 
   845 lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}"
   846   by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp)
   847 
   848 lemma finite_greaterThanAtMost_int [iff]: "finite {l<..(u::int)}"
   849   by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
   850 
   851 lemma finite_greaterThanLessThan_int [iff]: "finite {l<..<u::int}"
   852   by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
   853 
   854 
   855 subsubsection {* Cardinality *}
   856 
   857 lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u"
   858   apply (case_tac "0 \<le> u")
   859   apply (subst image_atLeastZeroLessThan_int, assumption)
   860   apply (subst card_image)
   861   apply (auto simp add: inj_on_def)
   862   done
   863 
   864 lemma card_atLeastLessThan_int [simp]: "card {l..<u} = nat (u - l)"
   865   apply (subgoal_tac "card {l..<u} = card {0..<u-l}")
   866   apply (erule ssubst, rule card_atLeastZeroLessThan_int)
   867   apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
   868   apply (erule subst)
   869   apply (rule card_image)
   870   apply (simp add: inj_on_def)
   871   apply (rule image_add_int_atLeastLessThan)
   872   done
   873 
   874 lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)"
   875 apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
   876 apply (auto simp add: algebra_simps)
   877 done
   878 
   879 lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)"
   880 by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
   881 
   882 lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))"
   883 by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
   884 
   885 lemma finite_M_bounded_by_nat: "finite {k. P k \<and> k < (i::nat)}"
   886 proof -
   887   have "{k. P k \<and> k < i} \<subseteq> {..<i}" by auto
   888   with finite_lessThan[of "i"] show ?thesis by (simp add: finite_subset)
   889 qed
   890 
   891 lemma card_less:
   892 assumes zero_in_M: "0 \<in> M"
   893 shows "card {k \<in> M. k < Suc i} \<noteq> 0"
   894 proof -
   895   from zero_in_M have "{k \<in> M. k < Suc i} \<noteq> {}" by auto
   896   with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff)
   897 qed
   898 
   899 lemma card_less_Suc2: "0 \<notin> M \<Longrightarrow> card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
   900 apply (rule card_bij_eq [of Suc _ _ "\<lambda>x. x - Suc 0"])
   901 apply simp
   902 apply fastsimp
   903 apply auto
   904 apply (rule inj_on_diff_nat)
   905 apply auto
   906 apply (case_tac x)
   907 apply auto
   908 apply (case_tac xa)
   909 apply auto
   910 apply (case_tac xa)
   911 apply auto
   912 done
   913 
   914 lemma card_less_Suc:
   915   assumes zero_in_M: "0 \<in> M"
   916     shows "Suc (card {k. Suc k \<in> M \<and> k < i}) = card {k \<in> M. k < Suc i}"
   917 proof -
   918   from assms have a: "0 \<in> {k \<in> M. k < Suc i}" by simp
   919   hence c: "{k \<in> M. k < Suc i} = insert 0 ({k \<in> M. k < Suc i} - {0})"
   920     by (auto simp only: insert_Diff)
   921   have b: "{k \<in> M. k < Suc i} - {0} = {k \<in> M - {0}. k < Suc i}"  by auto
   922   from finite_M_bounded_by_nat[of "\<lambda>x. x \<in> M" "Suc i"] have "Suc (card {k. Suc k \<in> M \<and> k < i}) = card (insert 0 ({k \<in> M. k < Suc i} - {0}))"
   923     apply (subst card_insert)
   924     apply simp_all
   925     apply (subst b)
   926     apply (subst card_less_Suc2[symmetric])
   927     apply simp_all
   928     done
   929   with c show ?thesis by simp
   930 qed
   931 
   932 
   933 subsection {*Lemmas useful with the summation operator setsum*}
   934 
   935 text {* For examples, see Algebra/poly/UnivPoly2.thy *}
   936 
   937 subsubsection {* Disjoint Unions *}
   938 
   939 text {* Singletons and open intervals *}
   940 
   941 lemma ivl_disj_un_singleton:
   942   "{l::'a::linorder} Un {l<..} = {l..}"
   943   "{..<u} Un {u::'a::linorder} = {..u}"
   944   "(l::'a::linorder) < u ==> {l} Un {l<..<u} = {l..<u}"
   945   "(l::'a::linorder) < u ==> {l<..<u} Un {u} = {l<..u}"
   946   "(l::'a::linorder) <= u ==> {l} Un {l<..u} = {l..u}"
   947   "(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}"
   948 by auto
   949 
   950 text {* One- and two-sided intervals *}
   951 
   952 lemma ivl_disj_un_one:
   953   "(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}"
   954   "(l::'a::linorder) <= u ==> {..<l} Un {l..<u} = {..<u}"
   955   "(l::'a::linorder) <= u ==> {..l} Un {l<..u} = {..u}"
   956   "(l::'a::linorder) <= u ==> {..<l} Un {l..u} = {..u}"
   957   "(l::'a::linorder) <= u ==> {l<..u} Un {u<..} = {l<..}"
   958   "(l::'a::linorder) < u ==> {l<..<u} Un {u..} = {l<..}"
   959   "(l::'a::linorder) <= u ==> {l..u} Un {u<..} = {l..}"
   960   "(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}"
   961 by auto
   962 
   963 text {* Two- and two-sided intervals *}
   964 
   965 lemma ivl_disj_un_two:
   966   "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
   967   "[| (l::'a::linorder) <= m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}"
   968   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..<u} = {l..<u}"
   969   "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}"
   970   "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..u} = {l<..u}"
   971   "[| (l::'a::linorder) <= m; m <= u |] ==> {l<..m} Un {m<..u} = {l<..u}"
   972   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..u} = {l..u}"
   973   "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m<..u} = {l..u}"
   974 by auto
   975 
   976 lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two
   977 
   978 subsubsection {* Disjoint Intersections *}
   979 
   980 text {* One- and two-sided intervals *}
   981 
   982 lemma ivl_disj_int_one:
   983   "{..l::'a::order} Int {l<..<u} = {}"
   984   "{..<l} Int {l..<u} = {}"
   985   "{..l} Int {l<..u} = {}"
   986   "{..<l} Int {l..u} = {}"
   987   "{l<..u} Int {u<..} = {}"
   988   "{l<..<u} Int {u..} = {}"
   989   "{l..u} Int {u<..} = {}"
   990   "{l..<u} Int {u..} = {}"
   991   by auto
   992 
   993 text {* Two- and two-sided intervals *}
   994 
   995 lemma ivl_disj_int_two:
   996   "{l::'a::order<..<m} Int {m..<u} = {}"
   997   "{l<..m} Int {m<..<u} = {}"
   998   "{l..<m} Int {m..<u} = {}"
   999   "{l..m} Int {m<..<u} = {}"
  1000   "{l<..<m} Int {m..u} = {}"
  1001   "{l<..m} Int {m<..u} = {}"
  1002   "{l..<m} Int {m..u} = {}"
  1003   "{l..m} Int {m<..u} = {}"
  1004   by auto
  1005 
  1006 lemmas ivl_disj_int = ivl_disj_int_one ivl_disj_int_two
  1007 
  1008 subsubsection {* Some Differences *}
  1009 
  1010 lemma ivl_diff[simp]:
  1011  "i \<le> n \<Longrightarrow> {i..<m} - {i..<n} = {n..<(m::'a::linorder)}"
  1012 by(auto)
  1013 
  1014 
  1015 subsubsection {* Some Subset Conditions *}
  1016 
  1017 lemma ivl_subset [simp,no_atp]:
  1018  "({i..<j} \<subseteq> {m..<n}) = (j \<le> i | m \<le> i & j \<le> (n::'a::linorder))"
  1019 apply(auto simp:linorder_not_le)
  1020 apply(rule ccontr)
  1021 apply(insert linorder_le_less_linear[of i n])
  1022 apply(clarsimp simp:linorder_not_le)
  1023 apply(fastsimp)
  1024 done
  1025 
  1026 
  1027 subsection {* Summation indexed over intervals *}
  1028 
  1029 syntax
  1030   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
  1031   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10)
  1032   "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<_./ _)" [0,0,10] 10)
  1033   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<=_./ _)" [0,0,10] 10)
  1034 syntax (xsymbols)
  1035   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
  1036   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
  1037   "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
  1038   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
  1039 syntax (HTML output)
  1040   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
  1041   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
  1042   "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
  1043   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
  1044 syntax (latex_sum output)
  1045   "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1046  ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10)
  1047   "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1048  ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10)
  1049   "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1050  ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10)
  1051   "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1052  ("(3\<^raw:$\sum_{>_ \<le> _\<^raw:}$> _)" [0,0,10] 10)
  1053 
  1054 translations
  1055   "\<Sum>x=a..b. t" == "CONST setsum (%x. t) {a..b}"
  1056   "\<Sum>x=a..<b. t" == "CONST setsum (%x. t) {a..<b}"
  1057   "\<Sum>i\<le>n. t" == "CONST setsum (\<lambda>i. t) {..n}"
  1058   "\<Sum>i<n. t" == "CONST setsum (\<lambda>i. t) {..<n}"
  1059 
  1060 text{* The above introduces some pretty alternative syntaxes for
  1061 summation over intervals:
  1062 \begin{center}
  1063 \begin{tabular}{lll}
  1064 Old & New & \LaTeX\\
  1065 @{term[source]"\<Sum>x\<in>{a..b}. e"} & @{term"\<Sum>x=a..b. e"} & @{term[mode=latex_sum]"\<Sum>x=a..b. e"}\\
  1066 @{term[source]"\<Sum>x\<in>{a..<b}. e"} & @{term"\<Sum>x=a..<b. e"} & @{term[mode=latex_sum]"\<Sum>x=a..<b. e"}\\
  1067 @{term[source]"\<Sum>x\<in>{..b}. e"} & @{term"\<Sum>x\<le>b. e"} & @{term[mode=latex_sum]"\<Sum>x\<le>b. e"}\\
  1068 @{term[source]"\<Sum>x\<in>{..<b}. e"} & @{term"\<Sum>x<b. e"} & @{term[mode=latex_sum]"\<Sum>x<b. e"}
  1069 \end{tabular}
  1070 \end{center}
  1071 The left column shows the term before introduction of the new syntax,
  1072 the middle column shows the new (default) syntax, and the right column
  1073 shows a special syntax. The latter is only meaningful for latex output
  1074 and has to be activated explicitly by setting the print mode to
  1075 @{text latex_sum} (e.g.\ via @{text "mode = latex_sum"} in
  1076 antiquotations). It is not the default \LaTeX\ output because it only
  1077 works well with italic-style formulae, not tt-style.
  1078 
  1079 Note that for uniformity on @{typ nat} it is better to use
  1080 @{term"\<Sum>x::nat=0..<n. e"} rather than @{text"\<Sum>x<n. e"}: @{text setsum} may
  1081 not provide all lemmas available for @{term"{m..<n}"} also in the
  1082 special form for @{term"{..<n}"}. *}
  1083 
  1084 text{* This congruence rule should be used for sums over intervals as
  1085 the standard theorem @{text[source]setsum_cong} does not work well
  1086 with the simplifier who adds the unsimplified premise @{term"x:B"} to
  1087 the context. *}
  1088 
  1089 lemma setsum_ivl_cong:
  1090  "\<lbrakk>a = c; b = d; !!x. \<lbrakk> c \<le> x; x < d \<rbrakk> \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow>
  1091  setsum f {a..<b} = setsum g {c..<d}"
  1092 by(rule setsum_cong, simp_all)
  1093 
  1094 (* FIXME why are the following simp rules but the corresponding eqns
  1095 on intervals are not? *)
  1096 
  1097 lemma setsum_atMost_Suc[simp]: "(\<Sum>i \<le> Suc n. f i) = (\<Sum>i \<le> n. f i) + f(Suc n)"
  1098 by (simp add:atMost_Suc add_ac)
  1099 
  1100 lemma setsum_lessThan_Suc[simp]: "(\<Sum>i < Suc n. f i) = (\<Sum>i < n. f i) + f n"
  1101 by (simp add:lessThan_Suc add_ac)
  1102 
  1103 lemma setsum_cl_ivl_Suc[simp]:
  1104   "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))"
  1105 by (auto simp:add_ac atLeastAtMostSuc_conv)
  1106 
  1107 lemma setsum_op_ivl_Suc[simp]:
  1108   "setsum f {m..<Suc n} = (if n < m then 0 else setsum f {m..<n} + f(n))"
  1109 by (auto simp:add_ac atLeastLessThanSuc)
  1110 (*
  1111 lemma setsum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==>
  1112     (\<Sum>i=n..m+1. f i) = (\<Sum>i=n..m. f i) + f(m + 1)"
  1113 by (auto simp:add_ac atLeastAtMostSuc_conv)
  1114 *)
  1115 
  1116 lemma setsum_head:
  1117   fixes n :: nat
  1118   assumes mn: "m <= n" 
  1119   shows "(\<Sum>x\<in>{m..n}. P x) = P m + (\<Sum>x\<in>{m<..n}. P x)" (is "?lhs = ?rhs")
  1120 proof -
  1121   from mn
  1122   have "{m..n} = {m} \<union> {m<..n}"
  1123     by (auto intro: ivl_disj_un_singleton)
  1124   hence "?lhs = (\<Sum>x\<in>{m} \<union> {m<..n}. P x)"
  1125     by (simp add: atLeast0LessThan)
  1126   also have "\<dots> = ?rhs" by simp
  1127   finally show ?thesis .
  1128 qed
  1129 
  1130 lemma setsum_head_Suc:
  1131   "m \<le> n \<Longrightarrow> setsum f {m..n} = f m + setsum f {Suc m..n}"
  1132 by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost)
  1133 
  1134 lemma setsum_head_upt_Suc:
  1135   "m < n \<Longrightarrow> setsum f {m..<n} = f m + setsum f {Suc m..<n}"
  1136 apply(insert setsum_head_Suc[of m "n - Suc 0" f])
  1137 apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] algebra_simps)
  1138 done
  1139 
  1140 lemma setsum_ub_add_nat: assumes "(m::nat) \<le> n + 1"
  1141   shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}"
  1142 proof-
  1143   have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using `m \<le> n+1` by auto
  1144   thus ?thesis by (auto simp: ivl_disj_int setsum_Un_disjoint
  1145     atLeastSucAtMost_greaterThanAtMost)
  1146 qed
  1147 
  1148 lemma setsum_add_nat_ivl: "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
  1149   setsum f {m..<n} + setsum f {n..<p} = setsum f {m..<p::nat}"
  1150 by (simp add:setsum_Un_disjoint[symmetric] ivl_disj_int ivl_disj_un)
  1151 
  1152 lemma setsum_diff_nat_ivl:
  1153 fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
  1154 shows "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
  1155   setsum f {m..<p} - setsum f {m..<n} = setsum f {n..<p}"
  1156 using setsum_add_nat_ivl [of m n p f,symmetric]
  1157 apply (simp add: add_ac)
  1158 done
  1159 
  1160 lemma setsum_natinterval_difff:
  1161   fixes f:: "nat \<Rightarrow> ('a::ab_group_add)"
  1162   shows  "setsum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
  1163           (if m <= n then f m - f(n + 1) else 0)"
  1164 by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
  1165 
  1166 lemma setsum_restrict_set':
  1167   "finite A \<Longrightarrow> setsum f {x \<in> A. x \<in> B} = (\<Sum>x\<in>A. if x \<in> B then f x else 0)"
  1168   by (simp add: setsum_restrict_set [symmetric] Int_def)
  1169 
  1170 lemma setsum_restrict_set'':
  1171   "finite A \<Longrightarrow> setsum f {x \<in> A. P x} = (\<Sum>x\<in>A. if P x  then f x else 0)"
  1172   by (simp add: setsum_restrict_set' [of A f "{x. P x}", simplified mem_Collect_eq])
  1173 
  1174 lemma setsum_setsum_restrict:
  1175   "finite S \<Longrightarrow> finite T \<Longrightarrow>
  1176     setsum (\<lambda>x. setsum (\<lambda>y. f x y) {y. y \<in> T \<and> R x y}) S = setsum (\<lambda>y. setsum (\<lambda>x. f x y) {x. x \<in> S \<and> R x y}) T"
  1177   by (simp add: setsum_restrict_set'') (rule setsum_commute)
  1178 
  1179 lemma setsum_image_gen: assumes fS: "finite S"
  1180   shows "setsum g S = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
  1181 proof-
  1182   { fix x assume "x \<in> S" then have "{y. y\<in> f`S \<and> f x = y} = {f x}" by auto }
  1183   hence "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
  1184     by simp
  1185   also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
  1186     by (rule setsum_setsum_restrict[OF fS finite_imageI[OF fS]])
  1187   finally show ?thesis .
  1188 qed
  1189 
  1190 lemma setsum_le_included:
  1191   fixes f :: "'a \<Rightarrow> 'b::ordered_comm_monoid_add"
  1192   assumes "finite s" "finite t"
  1193   and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"
  1194   shows "setsum f s \<le> setsum g t"
  1195 proof -
  1196   have "setsum f s \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) s"
  1197   proof (rule setsum_mono)
  1198     fix y assume "y \<in> s"
  1199     with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto
  1200     with assms show "f y \<le> setsum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
  1201       using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro]
  1202       by (auto intro!: setsum_mono2)
  1203   qed
  1204   also have "... \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) (i ` t)"
  1205     using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg)
  1206   also have "... \<le> setsum g t"
  1207     using assms by (auto simp: setsum_image_gen[symmetric])
  1208   finally show ?thesis .
  1209 qed
  1210 
  1211 lemma setsum_multicount_gen:
  1212   assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
  1213   shows "setsum (\<lambda>i. (card {j\<in>t. R i j})) s = setsum k t" (is "?l = ?r")
  1214 proof-
  1215   have "?l = setsum (\<lambda>i. setsum (\<lambda>x.1) {j\<in>t. R i j}) s" by auto
  1216   also have "\<dots> = ?r" unfolding setsum_setsum_restrict[OF assms(1-2)]
  1217     using assms(3) by auto
  1218   finally show ?thesis .
  1219 qed
  1220 
  1221 lemma setsum_multicount:
  1222   assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
  1223   shows "setsum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
  1224 proof-
  1225   have "?l = setsum (\<lambda>i. k) T" by(rule setsum_multicount_gen)(auto simp:assms)
  1226   also have "\<dots> = ?r" by(simp add: mult_commute)
  1227   finally show ?thesis by auto
  1228 qed
  1229 
  1230 
  1231 subsection{* Shifting bounds *}
  1232 
  1233 lemma setsum_shift_bounds_nat_ivl:
  1234   "setsum f {m+k..<n+k} = setsum (%i. f(i + k)){m..<n::nat}"
  1235 by (induct "n", auto simp:atLeastLessThanSuc)
  1236 
  1237 lemma setsum_shift_bounds_cl_nat_ivl:
  1238   "setsum f {m+k..n+k} = setsum (%i. f(i + k)){m..n::nat}"
  1239 apply (insert setsum_reindex[OF inj_on_add_nat, where h=f and B = "{m..n}"])
  1240 apply (simp add:image_add_atLeastAtMost o_def)
  1241 done
  1242 
  1243 corollary setsum_shift_bounds_cl_Suc_ivl:
  1244   "setsum f {Suc m..Suc n} = setsum (%i. f(Suc i)){m..n}"
  1245 by (simp add:setsum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
  1246 
  1247 corollary setsum_shift_bounds_Suc_ivl:
  1248   "setsum f {Suc m..<Suc n} = setsum (%i. f(Suc i)){m..<n}"
  1249 by (simp add:setsum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
  1250 
  1251 lemma setsum_shift_lb_Suc0_0:
  1252   "f(0::nat) = (0::nat) \<Longrightarrow> setsum f {Suc 0..k} = setsum f {0..k}"
  1253 by(simp add:setsum_head_Suc)
  1254 
  1255 lemma setsum_shift_lb_Suc0_0_upt:
  1256   "f(0::nat) = 0 \<Longrightarrow> setsum f {Suc 0..<k} = setsum f {0..<k}"
  1257 apply(cases k)apply simp
  1258 apply(simp add:setsum_head_upt_Suc)
  1259 done
  1260 
  1261 subsection {* The formula for geometric sums *}
  1262 
  1263 lemma geometric_sum:
  1264   assumes "x \<noteq> 1"
  1265   shows "(\<Sum>i=0..<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::field)"
  1266 proof -
  1267   from assms obtain y where "y = x - 1" and "y \<noteq> 0" by simp_all
  1268   moreover have "(\<Sum>i=0..<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
  1269   proof (induct n)
  1270     case 0 then show ?case by simp
  1271   next
  1272     case (Suc n)
  1273     moreover with `y \<noteq> 0` have "(1 + y) ^ n = (y * inverse y) * (1 + y) ^ n" by simp 
  1274     ultimately show ?case by (simp add: field_simps divide_inverse)
  1275   qed
  1276   ultimately show ?thesis by simp
  1277 qed
  1278 
  1279 
  1280 subsection {* The formula for arithmetic sums *}
  1281 
  1282 lemma gauss_sum:
  1283   "((1::'a::comm_semiring_1) + 1)*(\<Sum>i\<in>{1..n}. of_nat i) =
  1284    of_nat n*((of_nat n)+1)"
  1285 proof (induct n)
  1286   case 0
  1287   show ?case by simp
  1288 next
  1289   case (Suc n)
  1290   then show ?case by (simp add: algebra_simps)
  1291 qed
  1292 
  1293 theorem arith_series_general:
  1294   "((1::'a::comm_semiring_1) + 1) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
  1295   of_nat n * (a + (a + of_nat(n - 1)*d))"
  1296 proof cases
  1297   assume ngt1: "n > 1"
  1298   let ?I = "\<lambda>i. of_nat i" and ?n = "of_nat n"
  1299   have
  1300     "(\<Sum>i\<in>{..<n}. a+?I i*d) =
  1301      ((\<Sum>i\<in>{..<n}. a) + (\<Sum>i\<in>{..<n}. ?I i*d))"
  1302     by (rule setsum_addf)
  1303   also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
  1304   also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
  1305     unfolding One_nat_def
  1306     by (simp add: setsum_right_distrib atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt mult_ac)
  1307   also have "(1+1)*\<dots> = (1+1)*?n*a + d*(1+1)*(\<Sum>i\<in>{1..<n}. ?I i)"
  1308     by (simp add: left_distrib right_distrib)
  1309   also from ngt1 have "{1..<n} = {1..n - 1}"
  1310     by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost)
  1311   also from ngt1
  1312   have "(1+1)*?n*a + d*(1+1)*(\<Sum>i\<in>{1..n - 1}. ?I i) = ((1+1)*?n*a + d*?I (n - 1)*?I n)"
  1313     by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def)
  1314        (simp add:  mult_ac trans [OF add_commute of_nat_Suc [symmetric]])
  1315   finally show ?thesis by (simp add: algebra_simps)
  1316 next
  1317   assume "\<not>(n > 1)"
  1318   hence "n = 1 \<or> n = 0" by auto
  1319   thus ?thesis by (auto simp: algebra_simps)
  1320 qed
  1321 
  1322 lemma arith_series_nat:
  1323   "Suc (Suc 0) * (\<Sum>i\<in>{..<n}. a+i*d) = n * (a + (a+(n - 1)*d))"
  1324 proof -
  1325   have
  1326     "((1::nat) + 1) * (\<Sum>i\<in>{..<n::nat}. a + of_nat(i)*d) =
  1327     of_nat(n) * (a + (a + of_nat(n - 1)*d))"
  1328     by (rule arith_series_general)
  1329   thus ?thesis
  1330     unfolding One_nat_def by auto
  1331 qed
  1332 
  1333 lemma arith_series_int:
  1334   "(2::int) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
  1335   of_nat n * (a + (a + of_nat(n - 1)*d))"
  1336 proof -
  1337   have
  1338     "((1::int) + 1) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
  1339     of_nat(n) * (a + (a + of_nat(n - 1)*d))"
  1340     by (rule arith_series_general)
  1341   thus ?thesis by simp
  1342 qed
  1343 
  1344 lemma sum_diff_distrib:
  1345   fixes P::"nat\<Rightarrow>nat"
  1346   shows
  1347   "\<forall>x. Q x \<le> P x  \<Longrightarrow>
  1348   (\<Sum>x<n. P x) - (\<Sum>x<n. Q x) = (\<Sum>x<n. P x - Q x)"
  1349 proof (induct n)
  1350   case 0 show ?case by simp
  1351 next
  1352   case (Suc n)
  1353 
  1354   let ?lhs = "(\<Sum>x<n. P x) - (\<Sum>x<n. Q x)"
  1355   let ?rhs = "\<Sum>x<n. P x - Q x"
  1356 
  1357   from Suc have "?lhs = ?rhs" by simp
  1358   moreover
  1359   from Suc have "?lhs + P n - Q n = ?rhs + (P n - Q n)" by simp
  1360   moreover
  1361   from Suc have
  1362     "(\<Sum>x<n. P x) + P n - ((\<Sum>x<n. Q x) + Q n) = ?rhs + (P n - Q n)"
  1363     by (subst diff_diff_left[symmetric],
  1364         subst diff_add_assoc2)
  1365        (auto simp: diff_add_assoc2 intro: setsum_mono)
  1366   ultimately
  1367   show ?case by simp
  1368 qed
  1369 
  1370 subsection {* Products indexed over intervals *}
  1371 
  1372 syntax
  1373   "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10)
  1374   "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10)
  1375   "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _<_./ _)" [0,0,10] 10)
  1376   "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _<=_./ _)" [0,0,10] 10)
  1377 syntax (xsymbols)
  1378   "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _.._./ _)" [0,0,0,10] 10)
  1379   "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _..<_./ _)" [0,0,0,10] 10)
  1380   "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_<_./ _)" [0,0,10] 10)
  1381   "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_\<le>_./ _)" [0,0,10] 10)
  1382 syntax (HTML output)
  1383   "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _.._./ _)" [0,0,0,10] 10)
  1384   "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _..<_./ _)" [0,0,0,10] 10)
  1385   "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_<_./ _)" [0,0,10] 10)
  1386   "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_\<le>_./ _)" [0,0,10] 10)
  1387 syntax (latex_prod output)
  1388   "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1389  ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10)
  1390   "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1391  ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10)
  1392   "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1393  ("(3\<^raw:$\prod_{>_ < _\<^raw:}$> _)" [0,0,10] 10)
  1394   "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
  1395  ("(3\<^raw:$\prod_{>_ \<le> _\<^raw:}$> _)" [0,0,10] 10)
  1396 
  1397 translations
  1398   "\<Prod>x=a..b. t" == "CONST setprod (%x. t) {a..b}"
  1399   "\<Prod>x=a..<b. t" == "CONST setprod (%x. t) {a..<b}"
  1400   "\<Prod>i\<le>n. t" == "CONST setprod (\<lambda>i. t) {..n}"
  1401   "\<Prod>i<n. t" == "CONST setprod (\<lambda>i. t) {..<n}"
  1402 
  1403 subsection {* Transfer setup *}
  1404 
  1405 lemma transfer_nat_int_set_functions:
  1406     "{..n} = nat ` {0..int n}"
  1407     "{m..n} = nat ` {int m..int n}"  (* need all variants of these! *)
  1408   apply (auto simp add: image_def)
  1409   apply (rule_tac x = "int x" in bexI)
  1410   apply auto
  1411   apply (rule_tac x = "int x" in bexI)
  1412   apply auto
  1413   done
  1414 
  1415 lemma transfer_nat_int_set_function_closures:
  1416     "x >= 0 \<Longrightarrow> nat_set {x..y}"
  1417   by (simp add: nat_set_def)
  1418 
  1419 declare transfer_morphism_nat_int[transfer add
  1420   return: transfer_nat_int_set_functions
  1421     transfer_nat_int_set_function_closures
  1422 ]
  1423 
  1424 lemma transfer_int_nat_set_functions:
  1425     "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
  1426   by (simp only: is_nat_def transfer_nat_int_set_functions
  1427     transfer_nat_int_set_function_closures
  1428     transfer_nat_int_set_return_embed nat_0_le
  1429     cong: transfer_nat_int_set_cong)
  1430 
  1431 lemma transfer_int_nat_set_function_closures:
  1432     "is_nat x \<Longrightarrow> nat_set {x..y}"
  1433   by (simp only: transfer_nat_int_set_function_closures is_nat_def)
  1434 
  1435 declare transfer_morphism_int_nat[transfer add
  1436   return: transfer_int_nat_set_functions
  1437     transfer_int_nat_set_function_closures
  1438 ]
  1439 
  1440 end