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