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