src/HOL/Set_Interval.thy
author wenzelm
Thu, 10 Apr 2014 11:06:45 +0200
changeset 56507 5f6f2576a836
parent 56480 093ea91498e6
child 56949 d1a937cbf858
permissions -rw-r--r--
more contributors;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47317
432b29a96f61 modernized obsolete old-style theory name with proper new-style underscore
huffman
parents: 47222
diff changeset
     1
(*  Title:      HOL/Set_Interval.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
     2
    Author:     Tobias Nipkow
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
     3
    Author:     Clemens Ballarin
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
     4
    Author:     Jeremy Avigad
8924
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
     5
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
     6
lessThan, greaterThan, atLeast, atMost and two-sided intervals
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
     7
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
     8
Modern convention: Ixy stands for an interval where x and y
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
     9
describe the lower and upper bound and x,y : {c,o,i}
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
    10
where c = closed, o = open, i = infinite.
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
    11
Examples: Ico = {_ ..< _} and Ici = {_ ..}
8924
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
    12
*)
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
    13
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
    14
header {* Set intervals *}
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
    15
47317
432b29a96f61 modernized obsolete old-style theory name with proper new-style underscore
huffman
parents: 47222
diff changeset
    16
theory Set_Interval
55088
57c82e01022b moved 'bacc' back to 'Enum' (cf. 744934b818c7) -- reduces baggage loaded by 'Hilbert_Choice'
blanchet
parents: 55085
diff changeset
    17
imports Lattices_Big Nat_Transfer
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15056
diff changeset
    18
begin
8924
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
    19
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    20
context ord
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    21
begin
44008
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
    22
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    23
definition
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
    24
  lessThan    :: "'a => 'a set" ("(1{..<_})") where
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    25
  "{..<u} == {x. x < u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    26
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    27
definition
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
    28
  atMost      :: "'a => 'a set" ("(1{.._})") where
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    29
  "{..u} == {x. x \<le> u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    30
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    31
definition
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
    32
  greaterThan :: "'a => 'a set" ("(1{_<..})") where
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    33
  "{l<..} == {x. l<x}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    34
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    35
definition
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32596
diff changeset
    36
  atLeast     :: "'a => 'a set" ("(1{_..})") where
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    37
  "{l..} == {x. l\<le>x}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    38
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    39
definition
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    40
  greaterThanLessThan :: "'a => 'a => 'a set"  ("(1{_<..<_})") where
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    41
  "{l<..<u} == {l<..} Int {..<u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    42
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    43
definition
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    44
  atLeastLessThan :: "'a => 'a => 'a set"      ("(1{_..<_})") where
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    45
  "{l..<u} == {l..} Int {..<u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    46
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    47
definition
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    48
  greaterThanAtMost :: "'a => 'a => 'a set"    ("(1{_<.._})") where
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    49
  "{l<..u} == {l<..} Int {..u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    50
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    51
definition
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    52
  atLeastAtMost :: "'a => 'a => 'a set"        ("(1{_.._})") where
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    53
  "{l..u} == {l..} Int {..u}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    54
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
    55
end
8924
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
    56
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    57
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
    58
text{* A note of warning when using @{term"{..<n}"} on type @{typ
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
    59
nat}: it is equivalent to @{term"{0::nat..<n}"} but some lemmas involving
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
    60
@{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *}
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
    61
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    62
syntax
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    63
  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3UN _<=_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    64
  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3UN _<_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    65
  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3INT _<=_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    66
  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3INT _<_./ _)" [0, 0, 10] 10)
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    67
30372
96d508968153 UN syntax fix
nipkow
parents: 30242
diff changeset
    68
syntax (xsymbols)
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    69
  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _\<le>_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    70
  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _<_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    71
  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _\<le>_./ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    72
  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _<_./ _)" [0, 0, 10] 10)
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    73
30372
96d508968153 UN syntax fix
nipkow
parents: 30242
diff changeset
    74
syntax (latex output)
36364
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    75
  "_UNION_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ \<le> _)/ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    76
  "_UNION_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ < _)/ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    77
  "_INTER_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ \<le> _)/ _)" [0, 0, 10] 10)
0e2679025aeb fix syntax precedence declarations for UNION, INTER, SUP, INF
huffman
parents: 36307
diff changeset
    78
  "_INTER_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ < _)/ _)" [0, 0, 10] 10)
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    79
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    80
translations
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    81
  "UN i<=n. A"  == "UN i:{..n}. A"
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
    82
  "UN i<n. A"   == "UN i:{..<n}. A"
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    83
  "INT i<=n. A" == "INT i:{..n}. A"
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
    84
  "INT i<n. A"  == "INT i:{..<n}. A"
14418
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    85
b62323c85134 union/intersection over intervals
kleing
parents: 14398
diff changeset
    86
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
    87
subsection {* Various equivalences *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    88
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
    89
lemma (in ord) lessThan_iff [iff]: "(i: lessThan k) = (i<k)"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
    90
by (simp add: lessThan_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    91
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
    92
lemma Compl_lessThan [simp]:
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    93
    "!!k:: 'a::linorder. -lessThan k = atLeast k"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
    94
apply (auto simp add: lessThan_def atLeast_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    95
done
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    96
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
    97
lemma single_Diff_lessThan [simp]: "!!k:: 'a::order. {k} - lessThan k = {k}"
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
    98
by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
    99
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   100
lemma (in ord) greaterThan_iff [iff]: "(i: greaterThan k) = (k<i)"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   101
by (simp add: greaterThan_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   102
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   103
lemma Compl_greaterThan [simp]:
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   104
    "!!k:: 'a::linorder. -greaterThan k = atMost k"
26072
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25919
diff changeset
   105
  by (auto simp add: greaterThan_def atMost_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   106
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   107
lemma Compl_atMost [simp]: "!!k:: 'a::linorder. -atMost k = greaterThan k"
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   108
apply (subst Compl_greaterThan [symmetric])
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   109
apply (rule double_complement)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   110
done
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   111
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   112
lemma (in ord) atLeast_iff [iff]: "(i: atLeast k) = (k<=i)"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   113
by (simp add: atLeast_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   114
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   115
lemma Compl_atLeast [simp]:
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   116
    "!!k:: 'a::linorder. -atLeast k = lessThan k"
26072
f65a7fa2da6c <= and < on nat no longer depend on wellfounded relations
haftmann
parents: 25919
diff changeset
   117
  by (auto simp add: lessThan_def atLeast_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   118
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   119
lemma (in ord) atMost_iff [iff]: "(i: atMost k) = (i<=k)"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   120
by (simp add: atMost_def)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   121
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   122
lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   123
by (blast intro: order_antisym)
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   124
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   125
lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \<inter> { b <..} = { max a b <..}"
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   126
  by auto
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   127
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   128
lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \<inter> {..< b} = {..< min a b}"
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   129
  by auto
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   130
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   131
subsection {* Logical Equivalences for Set Inclusion and Equality *}
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   132
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   133
lemma atLeast_subset_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   134
     "(atLeast x \<subseteq> atLeast y) = (y \<le> (x::'a::order))"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   135
by (blast intro: order_trans)
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   136
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   137
lemma atLeast_eq_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   138
     "(atLeast x = atLeast y) = (x = (y::'a::linorder))"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   139
by (blast intro: order_antisym order_trans)
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   140
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   141
lemma greaterThan_subset_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   142
     "(greaterThan x \<subseteq> greaterThan y) = (y \<le> (x::'a::linorder))"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   143
apply (auto simp add: greaterThan_def)
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   144
 apply (subst linorder_not_less [symmetric], blast)
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   145
done
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   146
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   147
lemma greaterThan_eq_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   148
     "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   149
apply (rule iffI)
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   150
 apply (erule equalityE)
29709
cf8476cc440d fixed proposition slip
haftmann
parents: 29667
diff changeset
   151
 apply simp_all
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   152
done
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   153
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   154
lemma atMost_subset_iff [iff]: "(atMost x \<subseteq> atMost y) = (x \<le> (y::'a::order))"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   155
by (blast intro: order_trans)
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   156
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   157
lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::linorder))"
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   158
by (blast intro: order_antisym order_trans)
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   159
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   160
lemma lessThan_subset_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   161
     "(lessThan x \<subseteq> lessThan y) = (x \<le> (y::'a::linorder))"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   162
apply (auto simp add: lessThan_def)
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   163
 apply (subst linorder_not_less [symmetric], blast)
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   164
done
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   165
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   166
lemma lessThan_eq_iff [iff]:
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   167
     "(lessThan x = lessThan y) = (x = (y::'a::linorder))"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   168
apply (rule iffI)
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   169
 apply (erule equalityE)
29709
cf8476cc440d fixed proposition slip
haftmann
parents: 29667
diff changeset
   170
 apply simp_all
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   171
done
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   172
40703
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   173
lemma lessThan_strict_subset_iff:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   174
  fixes m n :: "'a::linorder"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   175
  shows "{..<m} < {..<n} \<longleftrightarrow> m < n"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   176
  by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq)
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   177
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
   178
subsection {*Two-sided intervals*}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   179
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   180
context ord
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   181
begin
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   182
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   183
lemma greaterThanLessThan_iff [simp]:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   184
  "(i : {l<..<u}) = (l < i & i < u)"
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   185
by (simp add: greaterThanLessThan_def)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   186
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   187
lemma atLeastLessThan_iff [simp]:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   188
  "(i : {l..<u}) = (l <= i & i < u)"
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   189
by (simp add: atLeastLessThan_def)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   190
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   191
lemma greaterThanAtMost_iff [simp]:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   192
  "(i : {l<..u}) = (l < i & i <= u)"
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   193
by (simp add: greaterThanAtMost_def)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   194
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
   195
lemma atLeastAtMost_iff [simp]:
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   196
  "(i : {l..u}) = (l <= i & i <= u)"
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   197
by (simp add: atLeastAtMost_def)
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   198
32436
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 32408
diff changeset
   199
text {* The above four lemmas could be declared as iffs. Unfortunately this
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52380
diff changeset
   200
breaks many proofs. Since it only helps blast, it is better to leave them
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52380
diff changeset
   201
alone. *}
32436
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 32408
diff changeset
   202
50999
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   203
lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \<inter> {..< b }"
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   204
  by auto
3de230ed0547 introduce order topology
hoelzl
parents: 50417
diff changeset
   205
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   206
end
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
   207
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   208
subsubsection{* Emptyness, singletons, subset *}
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   209
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   210
context order
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   211
begin
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   212
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   213
lemma atLeastatMost_empty[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   214
  "b < a \<Longrightarrow> {a..b} = {}"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   215
by(auto simp: atLeastAtMost_def atLeast_def atMost_def)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   216
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   217
lemma atLeastatMost_empty_iff[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   218
  "{a..b} = {} \<longleftrightarrow> (~ a <= b)"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   219
by auto (blast intro: order_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   220
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   221
lemma atLeastatMost_empty_iff2[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   222
  "{} = {a..b} \<longleftrightarrow> (~ a <= b)"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   223
by auto (blast intro: order_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   224
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   225
lemma atLeastLessThan_empty[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   226
  "b <= a \<Longrightarrow> {a..<b} = {}"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   227
by(auto simp: atLeastLessThan_def)
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   228
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   229
lemma atLeastLessThan_empty_iff[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   230
  "{a..<b} = {} \<longleftrightarrow> (~ a < b)"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   231
by auto (blast intro: le_less_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   232
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   233
lemma atLeastLessThan_empty_iff2[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   234
  "{} = {a..<b} \<longleftrightarrow> (~ a < b)"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   235
by auto (blast intro: le_less_trans)
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   236
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   237
lemma greaterThanAtMost_empty[simp]: "l \<le> k ==> {k<..l} = {}"
17719
2e75155c5ed5 Added a few lemmas
nipkow
parents: 17149
diff changeset
   238
by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def)
2e75155c5ed5 Added a few lemmas
nipkow
parents: 17149
diff changeset
   239
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   240
lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \<longleftrightarrow> ~ k < l"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   241
by auto (blast intro: less_le_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   242
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   243
lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \<longleftrightarrow> ~ k < l"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   244
by auto (blast intro: less_le_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   245
29709
cf8476cc440d fixed proposition slip
haftmann
parents: 29667
diff changeset
   246
lemma greaterThanLessThan_empty[simp]:"l \<le> k ==> {k<..<l} = {}"
17719
2e75155c5ed5 Added a few lemmas
nipkow
parents: 17149
diff changeset
   247
by(auto simp:greaterThanLessThan_def greaterThan_def lessThan_def)
2e75155c5ed5 Added a few lemmas
nipkow
parents: 17149
diff changeset
   248
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24853
diff changeset
   249
lemma atLeastAtMost_singleton [simp]: "{a..a} = {a}"
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   250
by (auto simp add: atLeastAtMost_def atMost_def atLeast_def)
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   251
36846
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   252
lemma atLeastAtMost_singleton': "a = b \<Longrightarrow> {a .. b} = {a}" by simp
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   253
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   254
lemma atLeastatMost_subset_iff[simp]:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   255
  "{a..b} <= {c..d} \<longleftrightarrow> (~ a <= b) | c <= a & b <= d"
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   256
unfolding atLeastAtMost_def atLeast_def atMost_def
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   257
by (blast intro: order_trans)
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   258
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   259
lemma atLeastatMost_psubset_iff:
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   260
  "{a..b} < {c..d} \<longleftrightarrow>
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   261
   ((~ a <= b) | c <= a & b <= d & (c < a | b < d))  &  c <= d"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   262
by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans)
32400
6c62363cf0d7 new lemmas
nipkow
parents: 32006
diff changeset
   263
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   264
lemma Icc_eq_Icc[simp]:
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   265
  "{l..h} = {l'..h'} = (l=l' \<and> h=h' \<or> \<not> l\<le>h \<and> \<not> l'\<le>h')"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   266
by(simp add: order_class.eq_iff)(auto intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   267
36846
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   268
lemma atLeastAtMost_singleton_iff[simp]:
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   269
  "{a .. b} = {c} \<longleftrightarrow> a = b \<and> b = c"
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   270
proof
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   271
  assume "{a..b} = {c}"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   272
  hence *: "\<not> (\<not> a \<le> b)" unfolding atLeastatMost_empty_iff[symmetric] by simp
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   273
  with `{a..b} = {c}` have "c \<le> a \<and> b \<le> c" by auto
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   274
  with * show "a = b \<and> b = c" by auto
36846
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   275
qed simp
0f67561ed5a6 Added atLeastAtMost_singleton_iff, atLeastAtMost_singleton'
hoelzl
parents: 36755
diff changeset
   276
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   277
lemma Icc_subset_Ici_iff[simp]:
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   278
  "{l..h} \<subseteq> {l'..} = (~ l\<le>h \<or> l\<ge>l')"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   279
by(auto simp: subset_eq intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   280
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   281
lemma Icc_subset_Iic_iff[simp]:
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   282
  "{l..h} \<subseteq> {..h'} = (~ l\<le>h \<or> h\<le>h')"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   283
by(auto simp: subset_eq intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   284
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   285
lemma not_Ici_eq_empty[simp]: "{l..} \<noteq> {}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   286
by(auto simp: set_eq_iff)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   287
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   288
lemma not_Iic_eq_empty[simp]: "{..h} \<noteq> {}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   289
by(auto simp: set_eq_iff)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   290
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   291
lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   292
lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   293
24691
e7f46ee04809 localized { .. } (but only a few thms)
nipkow
parents: 24449
diff changeset
   294
end
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   295
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   296
context no_top
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   297
begin
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   298
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   299
(* also holds for no_bot but no_top should suffice *)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   300
lemma not_UNIV_le_Icc[simp]: "\<not> UNIV \<subseteq> {l..h}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   301
using gt_ex[of h] by(auto simp: subset_eq less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   302
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   303
lemma not_UNIV_le_Iic[simp]: "\<not> UNIV \<subseteq> {..h}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   304
using gt_ex[of h] by(auto simp: subset_eq less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   305
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   306
lemma not_Ici_le_Icc[simp]: "\<not> {l..} \<subseteq> {l'..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   307
using gt_ex[of h']
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   308
by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   309
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   310
lemma not_Ici_le_Iic[simp]: "\<not> {l..} \<subseteq> {..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   311
using gt_ex[of h']
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   312
by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   313
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   314
end
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   315
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   316
context no_bot
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   317
begin
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   318
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   319
lemma not_UNIV_le_Ici[simp]: "\<not> UNIV \<subseteq> {l..}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   320
using lt_ex[of l] by(auto simp: subset_eq less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   321
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   322
lemma not_Iic_le_Icc[simp]: "\<not> {..h} \<subseteq> {l'..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   323
using lt_ex[of l']
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   324
by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   325
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   326
lemma not_Iic_le_Ici[simp]: "\<not> {..h} \<subseteq> {l'..}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   327
using lt_ex[of l']
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   328
by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   329
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   330
end
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   331
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   332
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   333
context no_top
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   334
begin
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   335
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   336
(* also holds for no_bot but no_top should suffice *)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   337
lemma not_UNIV_eq_Icc[simp]: "\<not> UNIV = {l'..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   338
using gt_ex[of h'] by(auto simp: set_eq_iff  less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   339
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   340
lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   341
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   342
lemma not_UNIV_eq_Iic[simp]: "\<not> UNIV = {..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   343
using gt_ex[of h'] by(auto simp: set_eq_iff  less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   344
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   345
lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   346
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   347
lemma not_Icc_eq_Ici[simp]: "\<not> {l..h} = {l'..}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   348
unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   349
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   350
lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   351
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   352
(* also holds for no_bot but no_top should suffice *)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   353
lemma not_Iic_eq_Ici[simp]: "\<not> {..h} = {l'..}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   354
using not_Ici_le_Iic[of l' h] by blast
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   355
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   356
lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   357
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   358
end
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   359
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   360
context no_bot
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   361
begin
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   362
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   363
lemma not_UNIV_eq_Ici[simp]: "\<not> UNIV = {l'..}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   364
using lt_ex[of l'] by(auto simp: set_eq_iff  less_le_not_le)
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   365
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   366
lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   367
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   368
lemma not_Icc_eq_Iic[simp]: "\<not> {l..h} = {..h'}"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   369
unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   370
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   371
lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric]
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   372
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   373
end
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   374
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   375
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 52729
diff changeset
   376
context dense_linorder
42891
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   377
begin
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   378
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   379
lemma greaterThanLessThan_empty_iff[simp]:
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   380
  "{ a <..< b } = {} \<longleftrightarrow> b \<le> a"
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   381
  using dense[of a b] by (cases "a < b") auto
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   382
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   383
lemma greaterThanLessThan_empty_iff2[simp]:
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   384
  "{} = { a <..< b } \<longleftrightarrow> b \<le> a"
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   385
  using dense[of a b] by (cases "a < b") auto
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   386
42901
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   387
lemma atLeastLessThan_subseteq_atLeastAtMost_iff:
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   388
  "{a ..< b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   389
  using dense[of "max a d" "b"]
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   390
  by (force simp: subset_eq Ball_def not_less[symmetric])
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   391
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   392
lemma greaterThanAtMost_subseteq_atLeastAtMost_iff:
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   393
  "{a <.. b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   394
  using dense[of "a" "min c b"]
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   395
  by (force simp: subset_eq Ball_def not_less[symmetric])
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   396
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   397
lemma greaterThanLessThan_subseteq_atLeastAtMost_iff:
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   398
  "{a <..< b} \<subseteq> { c .. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   399
  using dense[of "a" "min c b"] dense[of "max a d" "b"]
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   400
  by (force simp: subset_eq Ball_def not_less[symmetric])
e35cf2b25f48 equations for subsets of atLeastAtMost
hoelzl
parents: 42891
diff changeset
   401
43657
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   402
lemma atLeastAtMost_subseteq_atLeastLessThan_iff:
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   403
  "{a .. b} \<subseteq> { c ..< d } \<longleftrightarrow> (a \<le> b \<longrightarrow> c \<le> a \<and> b < d)"
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   404
  using dense[of "max a d" "b"]
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   405
  by (force simp: subset_eq Ball_def not_less[symmetric])
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   406
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   407
lemma greaterThanAtMost_subseteq_atLeastLessThan_iff:
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   408
  "{a <.. b} \<subseteq> { c ..< d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b < d)"
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   409
  using dense[of "a" "min c b"]
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   410
  by (force simp: subset_eq Ball_def not_less[symmetric])
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   411
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   412
lemma greaterThanLessThan_subseteq_atLeastLessThan_iff:
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   413
  "{a <..< b} \<subseteq> { c ..< d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   414
  using dense[of "a" "min c b"] dense[of "max a d" "b"]
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   415
  by (force simp: subset_eq Ball_def not_less[symmetric])
537ea3846f64 equalities of subsets of atLeastLessThan
hoelzl
parents: 43157
diff changeset
   416
56328
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   417
lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff:
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   418
  "{a <..< b} \<subseteq> { c <.. d } \<longleftrightarrow> (a < b \<longrightarrow> c \<le> a \<and> b \<le> d)"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   419
  using dense[of "a" "min c b"] dense[of "max a d" "b"]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   420
  by (force simp: subset_eq Ball_def not_less[symmetric])
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   421
42891
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   422
end
e2f473671937 simp rules for empty intervals on dense linear order
hoelzl
parents: 40703
diff changeset
   423
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   424
context no_top
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   425
begin
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   426
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   427
lemma greaterThan_non_empty[simp]: "{x <..} \<noteq> {}"
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   428
  using gt_ex[of x] by auto
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   429
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   430
end
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   431
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   432
context no_bot
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   433
begin
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   434
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   435
lemma lessThan_non_empty[simp]: "{..< x} \<noteq> {}"
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   436
  using lt_ex[of x] by auto
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   437
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   438
end
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   439
32408
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   440
lemma (in linorder) atLeastLessThan_subset_iff:
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   441
  "{a..<b} <= {c..<d} \<Longrightarrow> b <= a | c<=a & b<=d"
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   442
apply (auto simp:subset_eq Ball_def)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   443
apply(frule_tac x=a in spec)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   444
apply(erule_tac x=d in allE)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   445
apply (simp add: less_imp_le)
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   446
done
a1a85b0a26f7 new interval lemma
nipkow
parents: 32400
diff changeset
   447
40703
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   448
lemma atLeastLessThan_inj:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   449
  fixes a b c d :: "'a::linorder"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   450
  assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   451
  shows "a = c" "b = d"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   452
using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le linorder_antisym_conv2 subset_refl)+
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   453
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   454
lemma atLeastLessThan_eq_iff:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   455
  fixes a b c d :: "'a::linorder"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   456
  assumes "a < b" "c < d"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   457
  shows "{a ..< b} = {c ..< d} \<longleftrightarrow> a = c \<and> b = d"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   458
  using atLeastLessThan_inj assms by auto
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
   459
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52380
diff changeset
   460
lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \<longleftrightarrow> x = bot"
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   461
by (auto simp: set_eq_iff intro: le_bot)
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 51152
diff changeset
   462
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52380
diff changeset
   463
lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \<longleftrightarrow> x = top"
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   464
by (auto simp: set_eq_iff intro: top_le)
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 51152
diff changeset
   465
51334
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   466
lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff:
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   467
  "{x..y} = UNIV \<longleftrightarrow> (x = bot \<and> y = top)"
fd531bd984d8 more lemmas about intervals
nipkow
parents: 51329
diff changeset
   468
by (auto simp: set_eq_iff intro: top_le le_bot)
51328
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 51152
diff changeset
   469
d63ec23c9125 move auxiliary lemmas from Library/Extended_Reals to HOL image
hoelzl
parents: 51152
diff changeset
   470
56328
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   471
subsection {* Infinite intervals *}
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   472
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   473
context dense_linorder
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   474
begin
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   475
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   476
lemma infinite_Ioo:
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   477
  assumes "a < b"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   478
  shows "\<not> finite {a<..<b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   479
proof
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   480
  assume fin: "finite {a<..<b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   481
  moreover have ne: "{a<..<b} \<noteq> {}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   482
    using `a < b` by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   483
  ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   484
    using Max_in[of "{a <..< b}"] by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   485
  then obtain x where "Max {a <..< b} < x" "x < b"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   486
    using dense[of "Max {a<..<b}" b] by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   487
  then have "x \<in> {a <..< b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   488
    using `a < Max {a <..< b}` by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   489
  then have "x \<le> Max {a <..< b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   490
    using fin by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   491
  with `Max {a <..< b} < x` show False by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   492
qed
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   493
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   494
lemma infinite_Icc: "a < b \<Longrightarrow> \<not> finite {a .. b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   495
  using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   496
  by (auto dest: finite_subset)
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   497
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   498
lemma infinite_Ico: "a < b \<Longrightarrow> \<not> finite {a ..< b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   499
  using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   500
  by (auto dest: finite_subset)
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   501
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   502
lemma infinite_Ioc: "a < b \<Longrightarrow> \<not> finite {a <.. b}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   503
  using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   504
  by (auto dest: finite_subset)
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   505
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   506
end
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   507
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   508
lemma infinite_Iio: "\<not> finite {..< a :: 'a :: {no_bot, linorder}}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   509
proof
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   510
  assume "finite {..< a}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   511
  then have *: "\<And>x. x < a \<Longrightarrow> Min {..< a} \<le> x"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   512
    by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   513
  obtain x where "x < a"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   514
    using lt_ex by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   515
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   516
  obtain y where "y < Min {..< a}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   517
    using lt_ex by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   518
  also have "Min {..< a} \<le> x"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   519
    using `x < a` by fact
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   520
  also note `x < a`
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   521
  finally have "Min {..< a} \<le> y"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   522
    by fact
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   523
  with `y < Min {..< a}` show False by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   524
qed
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   525
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   526
lemma infinite_Iic: "\<not> finite {.. a :: 'a :: {no_bot, linorder}}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   527
  using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   528
  by (auto simp: subset_eq less_imp_le)
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   529
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   530
lemma infinite_Ioi: "\<not> finite {a :: 'a :: {no_top, linorder} <..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   531
proof
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   532
  assume "finite {a <..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   533
  then have *: "\<And>x. a < x \<Longrightarrow> x \<le> Max {a <..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   534
    by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   535
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   536
  obtain y where "Max {a <..} < y"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   537
    using gt_ex by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   538
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   539
  obtain x where "a < x"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   540
    using gt_ex by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   541
  also then have "x \<le> Max {a <..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   542
    by fact
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   543
  also note `Max {a <..} < y`
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   544
  finally have "y \<le> Max { a <..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   545
    by fact
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   546
  with `Max {a <..} < y` show False by auto
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   547
qed
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   548
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   549
lemma infinite_Ici: "\<not> finite {a :: 'a :: {no_top, linorder} ..}"
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   550
  using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"]
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   551
  by (auto simp: subset_eq less_imp_le)
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   552
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   553
subsubsection {* Intersection *}
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   554
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   555
context linorder
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   556
begin
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   557
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   558
lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   559
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   560
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   561
lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   562
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   563
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   564
lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   565
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   566
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   567
lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   568
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   569
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   570
lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   571
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   572
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   573
lemma Int_atLeastLessThan[simp]: "{a..<b} Int {c..<d} = {max a c ..< min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   574
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   575
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   576
lemma Int_greaterThanAtMost[simp]: "{a<..b} Int {c<..d} = {max a c <.. min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   577
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   578
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   579
lemma Int_greaterThanLessThan[simp]: "{a<..<b} Int {c<..<d} = {max a c <..< min b d}"
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   580
by auto
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   581
50417
f18b92f91818 add Int_atMost
hoelzl
parents: 47988
diff changeset
   582
lemma Int_atMost[simp]: "{..a} \<inter> {..b} = {.. min a b}"
f18b92f91818 add Int_atMost
hoelzl
parents: 47988
diff changeset
   583
  by (auto simp: min_def)
f18b92f91818 add Int_atMost
hoelzl
parents: 47988
diff changeset
   584
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   585
end
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   586
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   587
context complete_lattice
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   588
begin
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   589
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   590
lemma
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   591
  shows Sup_atLeast[simp]: "Sup {x ..} = top"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   592
    and Sup_greaterThanAtLeast[simp]: "x < top \<Longrightarrow> Sup {x <..} = top"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   593
    and Sup_atMost[simp]: "Sup {.. y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   594
    and Sup_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Sup { x .. y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   595
    and Sup_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Sup { x <.. y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   596
  by (auto intro!: Sup_eqI)
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   597
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   598
lemma
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   599
  shows Inf_atMost[simp]: "Inf {.. x} = bot"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   600
    and Inf_atMostLessThan[simp]: "top < x \<Longrightarrow> Inf {..< x} = bot"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   601
    and Inf_atLeast[simp]: "Inf {x ..} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   602
    and Inf_atLeastAtMost[simp]: "x \<le> y \<Longrightarrow> Inf { x .. y} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   603
    and Inf_atLeastLessThan[simp]: "x < y \<Longrightarrow> Inf { x ..< y} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   604
  by (auto intro!: Inf_eqI)
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   605
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   606
end
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   607
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   608
lemma
53216
ad2e09c30aa8 renamed inner_dense_linorder to dense_linorder
hoelzl
parents: 52729
diff changeset
   609
  fixes x y :: "'a :: {complete_lattice, dense_linorder}"
51329
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   610
  shows Sup_lessThan[simp]: "Sup {..< y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   611
    and Sup_atLeastLessThan[simp]: "x < y \<Longrightarrow> Sup { x ..< y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   612
    and Sup_greaterThanLessThan[simp]: "x < y \<Longrightarrow> Sup { x <..< y} = y"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   613
    and Inf_greaterThan[simp]: "Inf {x <..} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   614
    and Inf_greaterThanAtMost[simp]: "x < y \<Longrightarrow> Inf { x <.. y} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   615
    and Inf_greaterThanLessThan[simp]: "x < y \<Longrightarrow> Inf { x <..< y} = x"
4a3c453f99a1 split dense into inner_dense_order and no_top/no_bot
hoelzl
parents: 51328
diff changeset
   616
  by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded)
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   617
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   618
subsection {* Intervals of natural numbers *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   619
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   620
subsubsection {* The Constant @{term lessThan} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   621
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   622
lemma lessThan_0 [simp]: "lessThan (0::nat) = {}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   623
by (simp add: lessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   624
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   625
lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   626
by (simp add: lessThan_def less_Suc_eq, blast)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   627
43156
04aaac80183f fixed typo
kleing
parents: 42901
diff changeset
   628
text {* The following proof is convenient in induction proofs where
39072
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   629
new elements get indices at the beginning. So it is used to transform
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   630
@{term "{..<Suc n}"} to @{term "0::nat"} and @{term "{..< n}"}. *}
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   631
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   632
lemma lessThan_Suc_eq_insert_0: "{..<Suc n} = insert 0 (Suc ` {..<n})"
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   633
proof safe
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   634
  fix x assume "x < Suc n" "x \<notin> Suc ` {..<n}"
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   635
  then have "x \<noteq> Suc (x - 1)" by auto
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   636
  with `x < Suc n` show "x = 0" by auto
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   637
qed
1030b1a166ef Add lessThan_Suc_eq_insert_0
hoelzl
parents: 37664
diff changeset
   638
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   639
lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   640
by (simp add: lessThan_def atMost_def less_Suc_eq_le)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   641
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   642
lemma UN_lessThan_UNIV: "(UN m::nat. lessThan m) = UNIV"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   643
by blast
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   644
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   645
subsubsection {* The Constant @{term greaterThan} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   646
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   647
lemma greaterThan_0 [simp]: "greaterThan 0 = range Suc"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   648
apply (simp add: greaterThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   649
apply (blast dest: gr0_conv_Suc [THEN iffD1])
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   650
done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   651
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   652
lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   653
apply (simp add: greaterThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   654
apply (auto elim: linorder_neqE)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   655
done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   656
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   657
lemma INT_greaterThan_UNIV: "(INT m::nat. greaterThan m) = {}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   658
by blast
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   659
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   660
subsubsection {* The Constant @{term atLeast} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   661
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   662
lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   663
by (unfold atLeast_def UNIV_def, simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   664
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   665
lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   666
apply (simp add: atLeast_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   667
apply (simp add: Suc_le_eq)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   668
apply (simp add: order_le_less, blast)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   669
done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   670
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   671
lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   672
  by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   673
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   674
lemma UN_atLeast_UNIV: "(UN m::nat. atLeast m) = UNIV"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   675
by blast
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   676
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   677
subsubsection {* The Constant @{term atMost} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   678
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   679
lemma atMost_0 [simp]: "atMost (0::nat) = {0}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   680
by (simp add: atMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   681
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   682
lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   683
apply (simp add: atMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   684
apply (simp add: less_Suc_eq order_le_less, blast)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   685
done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   686
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   687
lemma UN_atMost_UNIV: "(UN m::nat. atMost m) = UNIV"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   688
by blast
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   689
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   690
subsubsection {* The Constant @{term atLeastLessThan} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   691
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   692
text{*The orientation of the following 2 rules is tricky. The lhs is
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   693
defined in terms of the rhs.  Hence the chosen orientation makes sense
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   694
in this theory --- the reverse orientation complicates proofs (eg
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   695
nontermination). But outside, when the definition of the lhs is rarely
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   696
used, the opposite orientation seems preferable because it reduces a
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   697
specific concept to a more general one. *}
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   698
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   699
lemma atLeast0LessThan: "{0::nat..<n} = {..<n}"
15042
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
   700
by(simp add:lessThan_def atLeastLessThan_def)
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   701
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   702
lemma atLeast0AtMost: "{0..n::nat} = {..n}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   703
by(simp add:atMost_def atLeastAtMost_def)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   704
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31509
diff changeset
   705
declare atLeast0LessThan[symmetric, code_unfold]
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31509
diff changeset
   706
        atLeast0AtMost[symmetric, code_unfold]
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   707
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   708
lemma atLeastLessThan0: "{m..<0::nat} = {}"
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   709
by (simp add: atLeastLessThan_def)
24449
2f05cb7fed85 added (code) lemmas for setsum and foldl
nipkow
parents: 24286
diff changeset
   710
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   711
subsubsection {* Intervals of nats with @{term Suc} *}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   712
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   713
text{*Not a simprule because the RHS is too messy.*}
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   714
lemma atLeastLessThanSuc:
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   715
    "{m..<Suc n} = (if m \<le> n then insert n {m..<n} else {})"
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   716
by (auto simp add: atLeastLessThan_def)
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   717
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   718
lemma atLeastLessThan_singleton [simp]: "{m..<Suc m} = {m}"
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   719
by (auto simp add: atLeastLessThan_def)
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
   720
(*
15047
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   721
lemma atLeast_sum_LessThan [simp]: "{m + k..<k::nat} = {}"
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   722
by (induct k, simp_all add: atLeastLessThanSuc)
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   723
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   724
lemma atLeastSucLessThan [simp]: "{Suc n..<n} = {}"
fa62de5862b9 redefining sumr to be a translation to setsum
paulson
parents: 15045
diff changeset
   725
by (auto simp add: atLeastLessThan_def)
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
   726
*)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   727
lemma atLeastLessThanSuc_atLeastAtMost: "{l..<Suc u} = {l..u}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   728
  by (simp add: lessThan_Suc_atMost atLeastAtMost_def atLeastLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   729
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   730
lemma atLeastSucAtMost_greaterThanAtMost: "{Suc l..u} = {l<..u}"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   731
  by (simp add: atLeast_Suc_greaterThan atLeastAtMost_def
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   732
    greaterThanAtMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   733
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   734
lemma atLeastSucLessThan_greaterThanLessThan: "{Suc l..<u} = {l<..<u}"
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
   735
  by (simp add: atLeast_Suc_greaterThan atLeastLessThan_def
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   736
    greaterThanLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   737
15554
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   738
lemma atLeastAtMostSuc_conv: "m \<le> Suc n \<Longrightarrow> {m..Suc n} = insert (Suc n) {m..n}"
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   739
by (auto simp add: atLeastAtMost_def)
03d4347b071d integrated Jeremy's FiniteLib
nipkow
parents: 15542
diff changeset
   740
45932
6f08f8fe9752 add lemmas
noschinl
parents: 44890
diff changeset
   741
lemma atLeastAtMost_insertL: "m \<le> n \<Longrightarrow> insert m {Suc m..n} = {m ..n}"
6f08f8fe9752 add lemmas
noschinl
parents: 44890
diff changeset
   742
by auto
6f08f8fe9752 add lemmas
noschinl
parents: 44890
diff changeset
   743
43157
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   744
text {* The analogous result is useful on @{typ int}: *}
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   745
(* here, because we don't have an own int section *)
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   746
lemma atLeastAtMostPlus1_int_conv:
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   747
  "m <= 1+n \<Longrightarrow> {m..1+n} = insert (1+n) {m..n::int}"
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   748
  by (auto intro: set_eqI)
b505be6f029a atLeastAtMostSuc_conv on int
kleing
parents: 43156
diff changeset
   749
33044
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   750
lemma atLeastLessThan_add_Un: "i \<le> j \<Longrightarrow> {i..<j+k} = {i..<j} \<union> {j..<j+k::nat}"
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   751
  apply (induct k) 
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   752
  apply (simp_all add: atLeastLessThanSuc)   
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   753
  done
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   754
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   755
subsubsection {* Image *}
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   756
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   757
lemma image_add_atLeastAtMost:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   758
  "(%n::nat. n+k) ` {i..j} = {i+k..j+k}" (is "?A = ?B")
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   759
proof
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   760
  show "?A \<subseteq> ?B" by auto
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   761
next
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   762
  show "?B \<subseteq> ?A"
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   763
  proof
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   764
    fix n assume a: "n : ?B"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19538
diff changeset
   765
    hence "n - k : {i..j}" by auto
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   766
    moreover have "n = (n - k) + k" using a by auto
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   767
    ultimately show "n : ?A" by blast
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   768
  qed
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   769
qed
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   770
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   771
lemma image_add_atLeastLessThan:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   772
  "(%n::nat. n+k) ` {i..<j} = {i+k..<j+k}" (is "?A = ?B")
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   773
proof
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   774
  show "?A \<subseteq> ?B" by auto
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   775
next
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   776
  show "?B \<subseteq> ?A"
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   777
  proof
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   778
    fix n assume a: "n : ?B"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19538
diff changeset
   779
    hence "n - k : {i..<j}" by auto
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   780
    moreover have "n = (n - k) + k" using a by auto
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   781
    ultimately show "n : ?A" by blast
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   782
  qed
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   783
qed
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   784
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   785
corollary image_Suc_atLeastAtMost[simp]:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   786
  "Suc ` {i..j} = {Suc i..Suc j}"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
   787
using image_add_atLeastAtMost[where k="Suc 0"] by simp
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   788
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   789
corollary image_Suc_atLeastLessThan[simp]:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   790
  "Suc ` {i..<j} = {Suc i..<Suc j}"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
   791
using image_add_atLeastLessThan[where k="Suc 0"] by simp
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   792
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   793
lemma image_add_int_atLeastLessThan:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   794
    "(%x. x + (l::int)) ` {0..<u-l} = {l..<u}"
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   795
  apply (auto simp add: image_def)
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   796
  apply (rule_tac x = "x - l" in bexI)
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   797
  apply auto
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   798
  done
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   799
37664
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   800
lemma image_minus_const_atLeastLessThan_nat:
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   801
  fixes c :: nat
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   802
  shows "(\<lambda>i. i - c) ` {x ..< y} =
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   803
      (if c < y then {x - c ..< y - c} else if x < y then {0} else {})"
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   804
    (is "_ = ?right")
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   805
proof safe
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   806
  fix a assume a: "a \<in> ?right"
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   807
  show "a \<in> (\<lambda>i. i - c) ` {x ..< y}"
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   808
  proof cases
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   809
    assume "c < y" with a show ?thesis
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   810
      by (auto intro!: image_eqI[of _ _ "a + c"])
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   811
  next
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   812
    assume "\<not> c < y" with a show ?thesis
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   813
      by (auto intro!: image_eqI[of _ _ x] split: split_if_asm)
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   814
  qed
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   815
qed auto
2946b8f057df Instantiate product type as euclidean space.
hoelzl
parents: 37388
diff changeset
   816
51152
b52cc015429a added lemma
Andreas Lochbihler
parents: 50999
diff changeset
   817
lemma image_int_atLeastLessThan: "int ` {a..<b} = {int a..<int b}"
55143
04448228381d explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
wenzelm
parents: 55088
diff changeset
   818
  by (auto intro!: image_eqI [where x = "nat x" for x])
51152
b52cc015429a added lemma
Andreas Lochbihler
parents: 50999
diff changeset
   819
35580
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   820
context ordered_ab_group_add
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   821
begin
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   822
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   823
lemma
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   824
  fixes x :: 'a
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   825
  shows image_uminus_greaterThan[simp]: "uminus ` {x<..} = {..<-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   826
  and image_uminus_atLeast[simp]: "uminus ` {x..} = {..-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   827
proof safe
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   828
  fix y assume "y < -x"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   829
  hence *:  "x < -y" using neg_less_iff_less[of "-y" x] by simp
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   830
  have "- (-y) \<in> uminus ` {x<..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   831
    by (rule imageI) (simp add: *)
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   832
  thus "y \<in> uminus ` {x<..}" by simp
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   833
next
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   834
  fix y assume "y \<le> -x"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   835
  have "- (-y) \<in> uminus ` {x..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   836
    by (rule imageI) (insert `y \<le> -x`[THEN le_imp_neg_le], simp)
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   837
  thus "y \<in> uminus ` {x..}" by simp
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   838
qed simp_all
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   839
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   840
lemma
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   841
  fixes x :: 'a
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   842
  shows image_uminus_lessThan[simp]: "uminus ` {..<x} = {-x<..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   843
  and image_uminus_atMost[simp]: "uminus ` {..x} = {-x..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   844
proof -
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   845
  have "uminus ` {..<x} = uminus ` uminus ` {-x<..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   846
    and "uminus ` {..x} = uminus ` uminus ` {-x..}" by simp_all
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   847
  thus "uminus ` {..<x} = {-x<..}" and "uminus ` {..x} = {-x..}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   848
    by (simp_all add: image_image
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   849
        del: image_uminus_greaterThan image_uminus_atLeast)
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   850
qed
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   851
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   852
lemma
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   853
  fixes x :: 'a
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   854
  shows image_uminus_atLeastAtMost[simp]: "uminus ` {x..y} = {-y..-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   855
  and image_uminus_greaterThanAtMost[simp]: "uminus ` {x<..y} = {-y..<-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   856
  and image_uminus_atLeastLessThan[simp]: "uminus ` {x..<y} = {-y<..-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   857
  and image_uminus_greaterThanLessThan[simp]: "uminus ` {x<..<y} = {-y<..<-x}"
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   858
  by (simp_all add: atLeastAtMost_def greaterThanAtMost_def atLeastLessThan_def
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   859
      greaterThanLessThan_def image_Int[OF inj_uminus] Int_commute)
0f74806cab22 Rewrite rules for images of minus of intervals
hoelzl
parents: 35216
diff changeset
   860
end
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
   861
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   862
subsubsection {* Finiteness *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   863
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   864
lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..<k}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   865
  by (induct k) (simp_all add: lessThan_Suc)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   866
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   867
lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   868
  by (induct k) (simp_all add: atMost_Suc)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   869
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   870
lemma finite_greaterThanLessThan [iff]:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   871
  fixes l :: nat shows "finite {l<..<u}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   872
by (simp add: greaterThanLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   873
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   874
lemma finite_atLeastLessThan [iff]:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   875
  fixes l :: nat shows "finite {l..<u}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   876
by (simp add: atLeastLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   877
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   878
lemma finite_greaterThanAtMost [iff]:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   879
  fixes l :: nat shows "finite {l<..u}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   880
by (simp add: greaterThanAtMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   881
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   882
lemma finite_atLeastAtMost [iff]:
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   883
  fixes l :: nat shows "finite {l..u}"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   884
by (simp add: atLeastAtMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   885
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   886
text {* A bounded set of natural numbers is finite. *}
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   887
lemma bounded_nat_set_is_finite:
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   888
  "(ALL i:N. i < (n::nat)) ==> finite N"
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   889
apply (rule finite_subset)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   890
 apply (rule_tac [2] finite_lessThan, auto)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   891
done
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   892
31044
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   893
text {* A set of natural numbers is finite iff it is bounded. *}
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   894
lemma finite_nat_set_iff_bounded:
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   895
  "finite(N::nat set) = (EX m. ALL n:N. n<m)" (is "?F = ?B")
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   896
proof
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   897
  assume f:?F  show ?B
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   898
    using Max_ge[OF `?F`, simplified less_Suc_eq_le[symmetric]] by blast
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   899
next
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   900
  assume ?B show ?F using `?B` by(blast intro:bounded_nat_set_is_finite)
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   901
qed
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   902
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   903
lemma finite_nat_set_iff_bounded_le:
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   904
  "finite(N::nat set) = (EX m. ALL n:N. n<=m)"
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   905
apply(simp add:finite_nat_set_iff_bounded)
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   906
apply(blast dest:less_imp_le_nat le_imp_less_Suc)
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   907
done
6896c2498ac0 new lemmas
nipkow
parents: 31017
diff changeset
   908
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   909
lemma finite_less_ub:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   910
     "!!f::nat=>nat. (!!n. n \<le> f n) ==> finite {n. f n \<le> u}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
   911
by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   912
56328
b3551501424e add rules about infinity of intervals
hoelzl
parents: 56238
diff changeset
   913
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   914
text{* Any subset of an interval of natural numbers the size of the
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   915
subset is exactly that interval. *}
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   916
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   917
lemma subset_card_intvl_is_intvl:
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 54606
diff changeset
   918
  assumes "A \<subseteq> {k..<k + card A}"
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 54606
diff changeset
   919
  shows "A = {k..<k + card A}"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   920
proof (cases "finite A")
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   921
  case True
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   922
  from this and assms show ?thesis
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   923
  proof (induct A rule: finite_linorder_max_induct)
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   924
    case empty thus ?case by auto
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   925
  next
33434
e9de8d69c1b9 fixed order of parameters in induction rules
nipkow
parents: 33318
diff changeset
   926
    case (insert b A)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   927
    hence *: "b \<notin> A" by auto
55085
0e8e4dc55866 moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
blanchet
parents: 54606
diff changeset
   928
    with insert have "A <= {k..<k + card A}" and "b = k + card A"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   929
      by fastforce+
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   930
    with insert * show ?case by auto
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   931
  qed
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   932
next
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   933
  case False
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53216
diff changeset
   934
  with assms show ?thesis by simp
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   935
qed
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   936
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
   937
32596
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   938
subsubsection {* Proving Inclusions and Equalities between Unions *}
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   939
36755
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   940
lemma UN_le_eq_Un0:
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   941
  "(\<Union>i\<le>n::nat. M i) = (\<Union>i\<in>{1..n}. M i) \<union> M 0" (is "?A = ?B")
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   942
proof
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   943
  show "?A <= ?B"
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   944
  proof
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   945
    fix x assume "x : ?A"
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   946
    then obtain i where i: "i\<le>n" "x : M i" by auto
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   947
    show "x : ?B"
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   948
    proof(cases i)
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   949
      case 0 with i show ?thesis by simp
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   950
    next
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   951
      case (Suc j) with i show ?thesis by auto
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   952
    qed
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   953
  qed
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   954
next
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   955
  show "?B <= ?A" by auto
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   956
qed
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   957
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   958
lemma UN_le_add_shift:
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   959
  "(\<Union>i\<le>n::nat. M(i+k)) = (\<Union>i\<in>{k..n+k}. M i)" (is "?A = ?B")
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   960
proof
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44008
diff changeset
   961
  show "?A <= ?B" by fastforce
36755
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   962
next
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   963
  show "?B <= ?A"
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   964
  proof
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   965
    fix x assume "x : ?B"
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   966
    then obtain i where i: "i : {k..n+k}" "x : M(i)" by auto
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   967
    hence "i-k\<le>n & x : M((i-k)+k)" by auto
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   968
    thus "x : ?A" by blast
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   969
  qed
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   970
qed
d1b498f2f50b added lemmas
nipkow
parents: 36365
diff changeset
   971
32596
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   972
lemma UN_UN_finite_eq: "(\<Union>n::nat. \<Union>i\<in>{0..<n}. A i) = (\<Union>n. A n)"
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   973
  by (auto simp add: atLeast0LessThan) 
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   974
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   975
lemma UN_finite_subset: "(!!n::nat. (\<Union>i\<in>{0..<n}. A i) \<subseteq> C) \<Longrightarrow> (\<Union>n. A n) \<subseteq> C"
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   976
  by (subst UN_UN_finite_eq [symmetric]) blast
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   977
33044
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   978
lemma UN_finite2_subset: 
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   979
     "(!!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)"
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   980
  apply (rule UN_finite_subset)
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   981
  apply (subst UN_UN_finite_eq [symmetric, of B]) 
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   982
  apply blast
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   983
  done
32596
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   984
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   985
lemma UN_finite2_eq:
33044
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   986
  "(!!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)"
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   987
  apply (rule subset_antisym)
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   988
   apply (rule UN_finite2_subset, blast)
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   989
 apply (rule UN_finite2_subset [where k=k])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35171
diff changeset
   990
 apply (force simp add: atLeastLessThan_add_Un [of 0])
33044
fd0a9c794ec1 Some new lemmas concerning sets
paulson
parents: 32960
diff changeset
   991
 done
32596
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   992
bd68c04dace1 New theorems for proving equalities and inclusions involving unions
paulson
parents: 32456
diff changeset
   993
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   994
subsubsection {* Cardinality *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   995
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
   996
lemma card_lessThan [simp]: "card {..<u} = u"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15140
diff changeset
   997
  by (induct u, simp_all add: lessThan_Suc)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   998
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
   999
lemma card_atMost [simp]: "card {..u} = Suc u"
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1000
  by (simp add: lessThan_Suc_atMost [THEN sym])
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1001
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1002
lemma card_atLeastLessThan [simp]: "card {l..<u} = u - l"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1003
  apply (subgoal_tac "card {l..<u} = card {..<u-l}")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1004
  apply (erule ssubst, rule card_lessThan)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1005
  apply (subgoal_tac "(%x. x + l) ` {..<u-l} = {l..<u}")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1006
  apply (erule subst)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1007
  apply (rule card_image)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1008
  apply (simp add: inj_on_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1009
  apply (auto simp add: image_def atLeastLessThan_def lessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1010
  apply (rule_tac x = "x - l" in exI)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1011
  apply arith
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1012
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1013
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1014
lemma card_atLeastAtMost [simp]: "card {l..u} = Suc u - l"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1015
  by (subst atLeastLessThanSuc_atLeastAtMost [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1016
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1017
lemma card_greaterThanAtMost [simp]: "card {l<..u} = u - l"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1018
  by (subst atLeastSucAtMost_greaterThanAtMost [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1019
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1020
lemma card_greaterThanLessThan [simp]: "card {l<..<u} = u - Suc l"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1021
  by (subst atLeastSucLessThan_greaterThanLessThan [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1022
26105
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1023
lemma ex_bij_betw_nat_finite:
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1024
  "finite M \<Longrightarrow> \<exists>h. bij_betw h {0..<card M} M"
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1025
apply(drule finite_imp_nat_seg_image_inj_on)
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1026
apply(auto simp:atLeast0LessThan[symmetric] lessThan_def[symmetric] card_image bij_betw_def)
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1027
done
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1028
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1029
lemma ex_bij_betw_finite_nat:
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1030
  "finite M \<Longrightarrow> \<exists>h. bij_betw h M {0..<card M}"
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1031
by (blast dest: ex_bij_betw_nat_finite bij_betw_inv)
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1032
31438
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1033
lemma finite_same_card_bij:
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1034
  "finite A \<Longrightarrow> finite B \<Longrightarrow> card A = card B \<Longrightarrow> EX h. bij_betw h A B"
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1035
apply(drule ex_bij_betw_finite_nat)
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1036
apply(drule ex_bij_betw_nat_finite)
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1037
apply(auto intro!:bij_betw_trans)
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1038
done
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1039
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1040
lemma ex_bij_betw_nat_finite_1:
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1041
  "finite M \<Longrightarrow> \<exists>h. bij_betw h {1 .. card M} M"
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1042
by (rule finite_same_card_bij) auto
a1c4c1500abe A few finite lemmas
nipkow
parents: 31044
diff changeset
  1043
40703
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1044
lemma bij_betw_iff_card:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1045
  assumes FIN: "finite A" and FIN': "finite B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1046
  shows BIJ: "(\<exists>f. bij_betw f A B) \<longleftrightarrow> (card A = card B)"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1047
using assms
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1048
proof(auto simp add: bij_betw_same_card)
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1049
  assume *: "card A = card B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1050
  obtain f where "bij_betw f A {0 ..< card A}"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1051
  using FIN ex_bij_betw_finite_nat by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1052
  moreover obtain g where "bij_betw g {0 ..< card B} B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1053
  using FIN' ex_bij_betw_nat_finite by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1054
  ultimately have "bij_betw (g o f) A B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1055
  using * by (auto simp add: bij_betw_trans)
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1056
  thus "(\<exists>f. bij_betw f A B)" by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1057
qed
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1058
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1059
lemma inj_on_iff_card_le:
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1060
  assumes FIN: "finite A" and FIN': "finite B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1061
  shows "(\<exists>f. inj_on f A \<and> f ` A \<le> B) = (card A \<le> card B)"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1062
proof (safe intro!: card_inj_on_le)
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1063
  assume *: "card A \<le> card B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1064
  obtain f where 1: "inj_on f A" and 2: "f ` A = {0 ..< card A}"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1065
  using FIN ex_bij_betw_finite_nat unfolding bij_betw_def by force
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1066
  moreover obtain g where "inj_on g {0 ..< card B}" and 3: "g ` {0 ..< card B} = B"
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1067
  using FIN' ex_bij_betw_nat_finite unfolding bij_betw_def by force
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1068
  ultimately have "inj_on g (f ` A)" using subset_inj_on[of g _ "f ` A"] * by force
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1069
  hence "inj_on (g o f) A" using 1 comp_inj_on by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1070
  moreover
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1071
  {have "{0 ..< card A} \<le> {0 ..< card B}" using * by force
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1072
   with 2 have "f ` A  \<le> {0 ..< card B}" by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1073
   hence "(g o f) ` A \<le> B" unfolding comp_def using 3 by force
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1074
  }
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1075
  ultimately show "(\<exists>f. inj_on f A \<and> f ` A \<le> B)" by blast
d1fc454d6735 Move some missing lemmas from Andrei Popescus 'Ordinals and Cardinals' AFP entry to the HOL-image.
hoelzl
parents: 39302
diff changeset
  1076
qed (insert assms, auto)
26105
ae06618225ec moved bij_betw from Library/FuncSet to Fun, redistributed some lemmas, and
nipkow
parents: 26072
diff changeset
  1077
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1078
subsection {* Intervals of integers *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1079
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1080
lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..<u+1} = {l..(u::int)}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1081
  by (auto simp add: atLeastAtMost_def atLeastLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1082
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1083
lemma atLeastPlusOneAtMost_greaterThanAtMost_int: "{l+1..u} = {l<..(u::int)}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1084
  by (auto simp add: atLeastAtMost_def greaterThanAtMost_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1085
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1086
lemma atLeastPlusOneLessThan_greaterThanLessThan_int:
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1087
    "{l+1..<u} = {l<..<u::int}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1088
  by (auto simp add: atLeastLessThan_def greaterThanLessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1089
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1090
subsubsection {* Finiteness *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1091
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1092
lemma image_atLeastZeroLessThan_int: "0 \<le> u ==>
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1093
    {(0::int)..<u} = int ` {..<nat u}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1094
  apply (unfold image_def lessThan_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1095
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1096
  apply (rule_tac x = "nat x" in exI)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35171
diff changeset
  1097
  apply (auto simp add: zless_nat_eq_int_zless [THEN sym])
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1098
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1099
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1100
lemma finite_atLeastZeroLessThan_int: "finite {(0::int)..<u}"
47988
e4b69e10b990 tuned proofs;
wenzelm
parents: 47317
diff changeset
  1101
  apply (cases "0 \<le> u")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1102
  apply (subst image_atLeastZeroLessThan_int, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1103
  apply (rule finite_imageI)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1104
  apply auto
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1105
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1106
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1107
lemma finite_atLeastLessThan_int [iff]: "finite {l..<u::int}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1108
  apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1109
  apply (erule subst)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1110
  apply (rule finite_imageI)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1111
  apply (rule finite_atLeastZeroLessThan_int)
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1112
  apply (rule image_add_int_atLeastLessThan)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1113
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1114
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1115
lemma finite_atLeastAtMost_int [iff]: "finite {l..(u::int)}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1116
  by (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1117
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1118
lemma finite_greaterThanAtMost_int [iff]: "finite {l<..(u::int)}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1119
  by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1120
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1121
lemma finite_greaterThanLessThan_int [iff]: "finite {l<..<u::int}"
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1122
  by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1123
24853
aab5798e5a33 added lemmas
nipkow
parents: 24748
diff changeset
  1124
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1125
subsubsection {* Cardinality *}
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1126
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1127
lemma card_atLeastZeroLessThan_int: "card {(0::int)..<u} = nat u"
47988
e4b69e10b990 tuned proofs;
wenzelm
parents: 47317
diff changeset
  1128
  apply (cases "0 \<le> u")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1129
  apply (subst image_atLeastZeroLessThan_int, assumption)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1130
  apply (subst card_image)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1131
  apply (auto simp add: inj_on_def)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1132
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1133
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1134
lemma card_atLeastLessThan_int [simp]: "card {l..<u} = nat (u - l)"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1135
  apply (subgoal_tac "card {l..<u} = card {0..<u-l}")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1136
  apply (erule ssubst, rule card_atLeastZeroLessThan_int)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1137
  apply (subgoal_tac "(%x. x + l) ` {0..<u-l} = {l..<u}")
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1138
  apply (erule subst)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1139
  apply (rule card_image)
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1140
  apply (simp add: inj_on_def)
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1141
  apply (rule image_add_int_atLeastLessThan)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1142
  done
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1143
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1144
lemma card_atLeastAtMost_int [simp]: "card {l..u} = nat (u - l + 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1145
apply (subst atLeastLessThanPlusOne_atLeastAtMost_int [THEN sym])
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1146
apply (auto simp add: algebra_simps)
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1147
done
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1148
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1149
lemma card_greaterThanAtMost_int [simp]: "card {l<..u} = nat (u - l)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1150
by (subst atLeastPlusOneAtMost_greaterThanAtMost_int [THEN sym], simp)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1151
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1152
lemma card_greaterThanLessThan_int [simp]: "card {l<..<u} = nat (u - (l + 1))"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1153
by (subst atLeastPlusOneLessThan_greaterThanLessThan_int [THEN sym], simp)
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1154
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1155
lemma finite_M_bounded_by_nat: "finite {k. P k \<and> k < (i::nat)}"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1156
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1157
  have "{k. P k \<and> k < i} \<subseteq> {..<i}" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1158
  with finite_lessThan[of "i"] show ?thesis by (simp add: finite_subset)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1159
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1160
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1161
lemma card_less:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1162
assumes zero_in_M: "0 \<in> M"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1163
shows "card {k \<in> M. k < Suc i} \<noteq> 0"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1164
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1165
  from zero_in_M have "{k \<in> M. k < Suc i} \<noteq> {}" by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1166
  with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1167
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1168
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1169
lemma card_less_Suc2: "0 \<notin> M \<Longrightarrow> card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
37388
793618618f78 tuned quotes, antiquotations and whitespace
haftmann
parents: 36846
diff changeset
  1170
apply (rule card_bij_eq [of Suc _ _ "\<lambda>x. x - Suc 0"])
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1171
apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44008
diff changeset
  1172
apply fastforce
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1173
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1174
apply (rule inj_on_diff_nat)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1175
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1176
apply (case_tac x)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1177
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1178
apply (case_tac xa)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1179
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1180
apply (case_tac xa)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1181
apply auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1182
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1183
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1184
lemma card_less_Suc:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1185
  assumes zero_in_M: "0 \<in> M"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1186
    shows "Suc (card {k. Suc k \<in> M \<and> k < i}) = card {k \<in> M. k < Suc i}"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1187
proof -
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1188
  from assms have a: "0 \<in> {k \<in> M. k < Suc i}" by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1189
  hence c: "{k \<in> M. k < Suc i} = insert 0 ({k \<in> M. k < Suc i} - {0})"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1190
    by (auto simp only: insert_Diff)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1191
  have b: "{k \<in> M. k < Suc i} - {0} = {k \<in> M - {0}. k < Suc i}"  by auto
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1192
  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}))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1193
    apply (subst card_insert)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1194
    apply simp_all
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1195
    apply (subst b)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1196
    apply (subst card_less_Suc2[symmetric])
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1197
    apply simp_all
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1198
    done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1199
  with c show ?thesis by simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1200
qed
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents: 26105
diff changeset
  1201
14485
ea2707645af8 new material from Avigad
paulson
parents: 14478
diff changeset
  1202
13850
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
  1203
subsection {*Lemmas useful with the summation operator setsum*}
6d1bb3059818 new logical equivalences
paulson
parents: 13735
diff changeset
  1204
16102
c5f6726d9bb1 Locale expressions: rename with optional mixfix syntax.
ballarin
parents: 16052
diff changeset
  1205
text {* For examples, see Algebra/poly/UnivPoly2.thy *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1206
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1207
subsubsection {* Disjoint Unions *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1208
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1209
text {* Singletons and open intervals *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1210
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1211
lemma ivl_disj_un_singleton:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1212
  "{l::'a::linorder} Un {l<..} = {l..}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1213
  "{..<u} Un {u::'a::linorder} = {..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1214
  "(l::'a::linorder) < u ==> {l} Un {l<..<u} = {l..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1215
  "(l::'a::linorder) < u ==> {l<..<u} Un {u} = {l<..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1216
  "(l::'a::linorder) <= u ==> {l} Un {l<..u} = {l..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1217
  "(l::'a::linorder) <= u ==> {l..<u} Un {u} = {l..u}"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 13850
diff changeset
  1218
by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1219
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1220
text {* One- and two-sided intervals *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1221
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1222
lemma ivl_disj_un_one:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1223
  "(l::'a::linorder) < u ==> {..l} Un {l<..<u} = {..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1224
  "(l::'a::linorder) <= u ==> {..<l} Un {l..<u} = {..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1225
  "(l::'a::linorder) <= u ==> {..l} Un {l<..u} = {..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1226
  "(l::'a::linorder) <= u ==> {..<l} Un {l..u} = {..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1227
  "(l::'a::linorder) <= u ==> {l<..u} Un {u<..} = {l<..}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1228
  "(l::'a::linorder) < u ==> {l<..<u} Un {u..} = {l<..}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1229
  "(l::'a::linorder) <= u ==> {l..u} Un {u<..} = {l..}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1230
  "(l::'a::linorder) <= u ==> {l..<u} Un {u..} = {l..}"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 13850
diff changeset
  1231
by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1232
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1233
text {* Two- and two-sided intervals *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1234
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1235
lemma ivl_disj_un_two:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1236
  "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..<u} = {l<..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1237
  "[| (l::'a::linorder) <= m; m < u |] ==> {l<..m} Un {m<..<u} = {l<..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1238
  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..<u} = {l..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1239
  "[| (l::'a::linorder) <= m; m < u |] ==> {l..m} Un {m<..<u} = {l..<u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1240
  "[| (l::'a::linorder) < m; m <= u |] ==> {l<..<m} Un {m..u} = {l<..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1241
  "[| (l::'a::linorder) <= m; m <= u |] ==> {l<..m} Un {m<..u} = {l<..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1242
  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..<m} Un {m..u} = {l..u}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1243
  "[| (l::'a::linorder) <= m; m <= u |] ==> {l..m} Un {m<..u} = {l..u}"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 13850
diff changeset
  1244
by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1245
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1246
lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1247
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1248
subsubsection {* Disjoint Intersections *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1249
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1250
text {* One- and two-sided intervals *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1251
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1252
lemma ivl_disj_int_one:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1253
  "{..l::'a::order} Int {l<..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1254
  "{..<l} Int {l..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1255
  "{..l} Int {l<..u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1256
  "{..<l} Int {l..u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1257
  "{l<..u} Int {u<..} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1258
  "{l<..<u} Int {u..} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1259
  "{l..u} Int {u<..} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1260
  "{l..<u} Int {u..} = {}"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 13850
diff changeset
  1261
  by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1262
14577
dbb95b825244 tuned document;
wenzelm
parents: 14485
diff changeset
  1263
text {* Two- and two-sided intervals *}
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1264
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1265
lemma ivl_disj_int_two:
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1266
  "{l::'a::order<..<m} Int {m..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1267
  "{l<..m} Int {m<..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1268
  "{l..<m} Int {m..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1269
  "{l..m} Int {m<..<u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1270
  "{l<..<m} Int {m..u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1271
  "{l<..m} Int {m<..u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1272
  "{l..<m} Int {m..u} = {}"
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 15042
diff changeset
  1273
  "{l..m} Int {m<..u} = {}"
14398
c5c47703f763 Efficient, graph-based reasoner for linear and partial orders.
ballarin
parents: 13850
diff changeset
  1274
  by auto
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1275
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
  1276
lemmas ivl_disj_int = ivl_disj_int_one ivl_disj_int_two
13735
7de9342aca7a HOL-Algebra partially ported to Isar.
ballarin
parents: 11609
diff changeset
  1277
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1278
subsubsection {* Some Differences *}
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1279
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1280
lemma ivl_diff[simp]:
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1281
 "i \<le> n \<Longrightarrow> {i..<m} - {i..<n} = {n..<(m::'a::linorder)}"
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1282
by(auto)
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1283
56194
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1284
lemma (in linorder) lessThan_minus_lessThan [simp]:
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1285
  "{..< n} - {..< m} = {m ..< n}"
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1286
  by auto
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1287
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1288
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1289
subsubsection {* Some Subset Conditions *}
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1290
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 53374
diff changeset
  1291
lemma ivl_subset [simp]:
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1292
 "({i..<j} \<subseteq> {m..<n}) = (j \<le> i | m \<le> i & j \<le> (n::'a::linorder))"
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1293
apply(auto simp:linorder_not_le)
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1294
apply(rule ccontr)
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1295
apply(insert linorder_le_less_linear[of i n])
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1296
apply(clarsimp simp:linorder_not_le)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44008
diff changeset
  1297
apply(fastforce)
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1298
done
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1299
15041
a6b1f0cef7b3 Got rid of Summation and made it a translation into setsum instead.
nipkow
parents: 14846
diff changeset
  1300
15042
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1301
subsection {* Summation indexed over intervals *}
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1302
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1303
syntax
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1304
  "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
  1305
  "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10)
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1306
  "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<_./ _)" [0,0,10] 10)
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1307
  "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<=_./ _)" [0,0,10] 10)
15042
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1308
syntax (xsymbols)
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1309
  "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
  1310
  "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1311
  "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1312
  "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
15042
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1313
syntax (HTML output)
fa7d27ef7e59 added {0::nat..n(} = {..n(}
nipkow
parents: 15041
diff changeset
  1314
  "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
  1315
  "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1316
  "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1317
  "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1318
syntax (latex_sum output)
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1319
  "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1320
 ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10)
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1321
  "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1322
 ("(3\<^raw:$\sum_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10)
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1323
  "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1324
 ("(3\<^raw:$\sum_{>_ < _\<^raw:}$> _)" [0,0,10] 10)
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1325
  "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1326
 ("(3\<^raw:$\sum_{>_ \<le> _\<^raw:}$> _)" [0,0,10] 10)
15041
a6b1f0cef7b3 Got rid of Summation and made it a translation into setsum instead.
nipkow
parents: 14846
diff changeset
  1327
15048
11b4dce71d73 more syntax
nipkow
parents: 15047
diff changeset
  1328
translations
28853
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28068
diff changeset
  1329
  "\<Sum>x=a..b. t" == "CONST setsum (%x. t) {a..b}"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28068
diff changeset
  1330
  "\<Sum>x=a..<b. t" == "CONST setsum (%x. t) {a..<b}"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28068
diff changeset
  1331
  "\<Sum>i\<le>n. t" == "CONST setsum (\<lambda>i. t) {..n}"
69eb69659bf3 Added new fold operator and renamed the old oe to fold_image.
nipkow
parents: 28068
diff changeset
  1332
  "\<Sum>i<n. t" == "CONST setsum (\<lambda>i. t) {..<n}"
15041
a6b1f0cef7b3 Got rid of Summation and made it a translation into setsum instead.
nipkow
parents: 14846
diff changeset
  1333
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1334
text{* The above introduces some pretty alternative syntaxes for
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1335
summation over intervals:
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1336
\begin{center}
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1337
\begin{tabular}{lll}
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1338
Old & New & \LaTeX\\
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1339
@{term[source]"\<Sum>x\<in>{a..b}. e"} & @{term"\<Sum>x=a..b. e"} & @{term[mode=latex_sum]"\<Sum>x=a..b. e"}\\
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1340
@{term[source]"\<Sum>x\<in>{a..<b}. e"} & @{term"\<Sum>x=a..<b. e"} & @{term[mode=latex_sum]"\<Sum>x=a..<b. e"}\\
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1341
@{term[source]"\<Sum>x\<in>{..b}. e"} & @{term"\<Sum>x\<le>b. e"} & @{term[mode=latex_sum]"\<Sum>x\<le>b. e"}\\
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1342
@{term[source]"\<Sum>x\<in>{..<b}. e"} & @{term"\<Sum>x<b. e"} & @{term[mode=latex_sum]"\<Sum>x<b. e"}
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1343
\end{tabular}
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1344
\end{center}
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1345
The left column shows the term before introduction of the new syntax,
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1346
the middle column shows the new (default) syntax, and the right column
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1347
shows a special syntax. The latter is only meaningful for latex output
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1348
and has to be activated explicitly by setting the print mode to
21502
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 20217
diff changeset
  1349
@{text latex_sum} (e.g.\ via @{text "mode = latex_sum"} in
15056
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1350
antiquotations). It is not the default \LaTeX\ output because it only
b75073d90bff Fine-tuned sum syntax.
nipkow
parents: 15052
diff changeset
  1351
works well with italic-style formulae, not tt-style.
15052
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1352
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1353
Note that for uniformity on @{typ nat} it is better to use
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1354
@{term"\<Sum>x::nat=0..<n. e"} rather than @{text"\<Sum>x<n. e"}: @{text setsum} may
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1355
not provide all lemmas available for @{term"{m..<n}"} also in the
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1356
special form for @{term"{..<n}"}. *}
cc562a263609 Added nice latex syntax.
nipkow
parents: 15048
diff changeset
  1357
15542
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1358
text{* This congruence rule should be used for sums over intervals as
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1359
the standard theorem @{text[source]setsum_cong} does not work well
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1360
with the simplifier who adds the unsimplified premise @{term"x:B"} to
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1361
the context. *}
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1362
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1363
lemma setsum_ivl_cong:
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1364
 "\<lbrakk>a = c; b = d; !!x. \<lbrakk> c \<le> x; x < d \<rbrakk> \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow>
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1365
 setsum f {a..<b} = setsum g {c..<d}"
ee6cd48cf840 more fine tuniung
nipkow
parents: 15539
diff changeset
  1366
by(rule setsum_cong, simp_all)
15041
a6b1f0cef7b3 Got rid of Summation and made it a translation into setsum instead.
nipkow
parents: 14846
diff changeset
  1367
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1368
(* FIXME why are the following simp rules but the corresponding eqns
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1369
on intervals are not? *)
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1370
16052
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1371
lemma setsum_atMost_Suc[simp]: "(\<Sum>i \<le> Suc n. f i) = (\<Sum>i \<le> n. f i) + f(Suc n)"
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1372
by (simp add:atMost_Suc add_ac)
880b0e786c1b tuned setsum rewrites
nipkow
parents: 16041
diff changeset
  1373
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1374
lemma setsum_lessThan_Suc[simp]: "(\<Sum>i < Suc n. f i) = (\<Sum>i < n. f i) + f n"
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1375
by (simp add:lessThan_Suc add_ac)
15041
a6b1f0cef7b3 Got rid of Summation and made it a translation into setsum instead.
nipkow
parents: 14846
diff changeset
  1376
15911
b730b0edc085 turned 2 lemmas into simp rules
nipkow
parents: 15561
diff changeset
  1377
lemma setsum_cl_ivl_Suc[simp]:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1378
  "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))"
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1379
by (auto simp:add_ac atLeastAtMostSuc_conv)
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1380
15911
b730b0edc085 turned 2 lemmas into simp rules
nipkow
parents: 15561
diff changeset
  1381
lemma setsum_op_ivl_Suc[simp]:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1382
  "setsum f {m..<Suc n} = (if n < m then 0 else setsum f {m..<n} + f(n))"
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1383
by (auto simp:add_ac atLeastLessThanSuc)
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1384
(*
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1385
lemma setsum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==>
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1386
    (\<Sum>i=n..m+1. f i) = (\<Sum>i=n..m. f i) + f(m + 1)"
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15554
diff changeset
  1387
by (auto simp:add_ac atLeastAtMostSuc_conv)
16041
5a8736668ced simplifier trace info; Suc-intervals
nipkow
parents: 15911
diff changeset
  1388
*)
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1389
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1390
lemma setsum_head:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1391
  fixes n :: nat
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1392
  assumes mn: "m <= n" 
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1393
  shows "(\<Sum>x\<in>{m..n}. P x) = P m + (\<Sum>x\<in>{m<..n}. P x)" (is "?lhs = ?rhs")
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1394
proof -
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1395
  from mn
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1396
  have "{m..n} = {m} \<union> {m<..n}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1397
    by (auto intro: ivl_disj_un_singleton)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1398
  hence "?lhs = (\<Sum>x\<in>{m} \<union> {m<..n}. P x)"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1399
    by (simp add: atLeast0LessThan)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1400
  also have "\<dots> = ?rhs" by simp
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1401
  finally show ?thesis .
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1402
qed
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1403
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1404
lemma setsum_head_Suc:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1405
  "m \<le> n \<Longrightarrow> setsum f {m..n} = f m + setsum f {Suc m..n}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1406
by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1407
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1408
lemma setsum_head_upt_Suc:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1409
  "m < n \<Longrightarrow> setsum f {m..<n} = f m + setsum f {Suc m..<n}"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1410
apply(insert setsum_head_Suc[of m "n - Suc 0" f])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28853
diff changeset
  1411
apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] algebra_simps)
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1412
done
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1413
31501
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1414
lemma setsum_ub_add_nat: assumes "(m::nat) \<le> n + 1"
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1415
  shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}"
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1416
proof-
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1417
  have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using `m \<le> n+1` by auto
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1418
  thus ?thesis by (auto simp: ivl_disj_int setsum_Un_disjoint
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1419
    atLeastSucAtMost_greaterThanAtMost)
2a60c9b951e0 New lemma
nipkow
parents: 31438
diff changeset
  1420
qed
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1421
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1422
lemma setsum_add_nat_ivl: "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1423
  setsum f {m..<n} + setsum f {n..<p} = setsum f {m..<p::nat}"
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1424
by (simp add:setsum_Un_disjoint[symmetric] ivl_disj_int ivl_disj_un)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1425
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1426
lemma setsum_diff_nat_ivl:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1427
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1428
shows "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1429
  setsum f {m..<p} - setsum f {m..<n} = setsum f {n..<p}"
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1430
using setsum_add_nat_ivl [of m n p f,symmetric]
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1431
apply (simp add: add_ac)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1432
done
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1433
31505
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1434
lemma setsum_natinterval_difff:
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1435
  fixes f:: "nat \<Rightarrow> ('a::ab_group_add)"
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1436
  shows  "setsum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1437
          (if m <= n then f m - f(n + 1) else 0)"
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1438
by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
6f589131ba94 new lemma
nipkow
parents: 31501
diff changeset
  1439
44008
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1440
lemma setsum_restrict_set':
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1441
  "finite A \<Longrightarrow> setsum f {x \<in> A. x \<in> B} = (\<Sum>x\<in>A. if x \<in> B then f x else 0)"
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1442
  by (simp add: setsum_restrict_set [symmetric] Int_def)
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1443
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1444
lemma setsum_restrict_set'':
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1445
  "finite A \<Longrightarrow> setsum f {x \<in> A. P x} = (\<Sum>x\<in>A. if P x  then f x else 0)"
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1446
  by (simp add: setsum_restrict_set' [of A f "{x. P x}", simplified mem_Collect_eq])
31509
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1447
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1448
lemma setsum_setsum_restrict:
44008
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1449
  "finite S \<Longrightarrow> finite T \<Longrightarrow>
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1450
    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"
2e09299ce807 tuned proofs
haftmann
parents: 43657
diff changeset
  1451
  by (simp add: setsum_restrict_set'') (rule setsum_commute)
31509
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1452
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1453
lemma setsum_image_gen: assumes fS: "finite S"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1454
  shows "setsum g S = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1455
proof-
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1456
  { fix x assume "x \<in> S" then have "{y. y\<in> f`S \<and> f x = y} = {f x}" by auto }
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1457
  hence "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1458
    by simp
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1459
  also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1460
    by (rule setsum_setsum_restrict[OF fS finite_imageI[OF fS]])
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1461
  finally show ?thesis .
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1462
qed
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1463
35171
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1464
lemma setsum_le_included:
36307
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1465
  fixes f :: "'a \<Rightarrow> 'b::ordered_comm_monoid_add"
35171
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1466
  assumes "finite s" "finite t"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1467
  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)"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1468
  shows "setsum f s \<le> setsum g t"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1469
proof -
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1470
  have "setsum f s \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) s"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1471
  proof (rule setsum_mono)
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1472
    fix y assume "y \<in> s"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1473
    with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1474
    with assms show "f y \<le> setsum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1475
      using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro]
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1476
      by (auto intro!: setsum_mono2)
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1477
  qed
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1478
  also have "... \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) (i ` t)"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1479
    using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg)
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1480
  also have "... \<le> setsum g t"
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1481
    using assms by (auto simp: setsum_image_gen[symmetric])
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1482
  finally show ?thesis .
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1483
qed
28f824c7addc Moved setprod_mono, abs_setprod and setsum_le_included to the Main image. Is used in Multivariate_Analysis.
hoelzl
parents: 35115
diff changeset
  1484
31509
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1485
lemma setsum_multicount_gen:
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1486
  assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1487
  shows "setsum (\<lambda>i. (card {j\<in>t. R i j})) s = setsum k t" (is "?l = ?r")
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1488
proof-
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1489
  have "?l = setsum (\<lambda>i. setsum (\<lambda>x.1) {j\<in>t. R i j}) s" by auto
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1490
  also have "\<dots> = ?r" unfolding setsum_setsum_restrict[OF assms(1-2)]
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1491
    using assms(3) by auto
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1492
  finally show ?thesis .
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1493
qed
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1494
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1495
lemma setsum_multicount:
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1496
  assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1497
  shows "setsum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1498
proof-
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1499
  have "?l = setsum (\<lambda>i. k) T" by(rule setsum_multicount_gen)(auto simp:assms)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35171
diff changeset
  1500
  also have "\<dots> = ?r" by(simp add: mult_commute)
31509
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1501
  finally show ?thesis by auto
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1502
qed
00ede188c5d6 more lemmas
nipkow
parents: 31505
diff changeset
  1503
56194
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1504
lemma setsum_nat_group: "(\<Sum>m<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {..< n * k}"
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1505
  apply (subgoal_tac "k = 0 | 0 < k", auto)
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1506
  apply (induct "n")
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1507
  apply (simp_all add: setsum_add_nat_ivl add_commute atLeast0LessThan[symmetric])
9ffbb4004c81 fix HOL-NSA; move lemmas
hoelzl
parents: 56193
diff changeset
  1508
  done
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1509
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1510
subsection{* Shifting bounds *}
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1511
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1512
lemma setsum_shift_bounds_nat_ivl:
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1513
  "setsum f {m+k..<n+k} = setsum (%i. f(i + k)){m..<n::nat}"
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1514
by (induct "n", auto simp:atLeastLessThanSuc)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15418
diff changeset
  1515
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1516
lemma setsum_shift_bounds_cl_nat_ivl:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1517
  "setsum f {m+k..n+k} = setsum (%i. f(i + k)){m..n::nat}"
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1518
apply (insert setsum_reindex[OF inj_on_add_nat, where h=f and B = "{m..n}"])
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1519
apply (simp add:image_add_atLeastAtMost o_def)
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1520
done
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1521
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1522
corollary setsum_shift_bounds_cl_Suc_ivl:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1523
  "setsum f {Suc m..Suc n} = setsum (%i. f(Suc i)){m..n}"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1524
by (simp add:setsum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1525
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1526
corollary setsum_shift_bounds_Suc_ivl:
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1527
  "setsum f {Suc m..<Suc n} = setsum (%i. f(Suc i)){m..<n}"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1528
by (simp add:setsum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
16733
236dfafbeb63 linear arithmetic now takes "&" in assumptions apart.
nipkow
parents: 16102
diff changeset
  1529
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1530
lemma setsum_shift_lb_Suc0_0:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1531
  "f(0::nat) = (0::nat) \<Longrightarrow> setsum f {Suc 0..k} = setsum f {0..k}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1532
by(simp add:setsum_head_Suc)
19106
6e6b5b1fdc06 * added Library/ASeries (sum of arithmetic series with instantiation to nat and int)
kleing
parents: 19022
diff changeset
  1533
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1534
lemma setsum_shift_lb_Suc0_0_upt:
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1535
  "f(0::nat) = 0 \<Longrightarrow> setsum f {Suc 0..<k} = setsum f {0..<k}"
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1536
apply(cases k)apply simp
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1537
apply(simp add:setsum_head_upt_Suc)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1538
done
19022
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1539
52380
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1540
lemma setsum_atMost_Suc_shift:
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1541
  fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1542
  shows "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1543
proof (induct n)
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1544
  case 0 show ?case by simp
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1545
next
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1546
  case (Suc n) note IH = this
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1547
  have "(\<Sum>i\<le>Suc (Suc n). f i) = (\<Sum>i\<le>Suc n. f i) + f (Suc (Suc n))"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1548
    by (rule setsum_atMost_Suc)
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1549
  also have "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1550
    by (rule IH)
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1551
  also have "f 0 + (\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) =
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1552
             f 0 + ((\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)))"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1553
    by (rule add_assoc)
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1554
  also have "(\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = (\<Sum>i\<le>Suc n. f (Suc i))"
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1555
    by (rule setsum_atMost_Suc [symmetric])
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1556
  finally show ?case .
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1557
qed
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1558
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1559
lemma setsum_last_plus: fixes n::nat shows "m <= n \<Longrightarrow> (\<Sum>i = m..n. f i) = f n + (\<Sum>i = m..<n. f i)"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1560
  by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost add_commute)
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1561
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1562
lemma setsum_Suc_diff:
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1563
  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1564
  assumes "m \<le> Suc n"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1565
  shows "(\<Sum>i = m..n. f(Suc i) - f i) = f (Suc n) - f m"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1566
using assms by (induct n) (auto simp: le_Suc_eq)
55718
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1567
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1568
lemma nested_setsum_swap:
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1569
     "(\<Sum>i = 0..n. (\<Sum>j = 0..<i. a i j)) = (\<Sum>j = 0..<n. \<Sum>i = Suc j..n. a i j)"
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1570
  by (induction n) (auto simp: setsum_addf)
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1571
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1572
lemma nested_setsum_swap':
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1573
     "(\<Sum>i\<le>n. (\<Sum>j<i. a i j)) = (\<Sum>j<n. \<Sum>i = Suc j..n. a i j)"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1574
  by (induction n) (auto simp: setsum_addf)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1575
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1576
lemma setsum_zero_power [simp]:
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1577
  fixes c :: "nat \<Rightarrow> 'a::division_ring"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1578
  shows "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1579
apply (cases "finite A")
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1580
  by (induction A rule: finite_induct) auto
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 56194
diff changeset
  1581
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1582
lemma setsum_zero_power' [simp]:
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1583
  fixes c :: "nat \<Rightarrow> 'a::field"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1584
  shows "(\<Sum>i\<in>A. c i * 0^i / d i) = (if finite A \<and> 0 \<in> A then c 0 / d 0 else 0)"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1585
  using setsum_zero_power [of "\<lambda>i. c i / d i" A]
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1586
  by auto
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56215
diff changeset
  1587
52380
3cc46b8cca5e lifting for primitive definitions;
haftmann
parents: 51334
diff changeset
  1588
17149
e2b19c92ef51 Lemmas on dvd, power and finite summation added or strengthened.
ballarin
parents: 16733
diff changeset
  1589
subsection {* The formula for geometric sums *}
e2b19c92ef51 Lemmas on dvd, power and finite summation added or strengthened.
ballarin
parents: 16733
diff changeset
  1590
e2b19c92ef51 Lemmas on dvd, power and finite summation added or strengthened.
ballarin
parents: 16733
diff changeset
  1591
lemma geometric_sum:
36307
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1592
  assumes "x \<noteq> 1"
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 55719
diff changeset
  1593
  shows "(\<Sum>i<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::field)"
36307
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1594
proof -
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1595
  from assms obtain y where "y = x - 1" and "y \<noteq> 0" by simp_all
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 55719
diff changeset
  1596
  moreover have "(\<Sum>i<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56328
diff changeset
  1597
    by (induct n) (simp_all add: field_simps `y \<noteq> 0`)
36307
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1598
  ultimately show ?thesis by simp
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1599
qed
1732232f9b27 sharpened constraint (c.f. 4e7f5b22dd7d); explicit is better than implicit
haftmann
parents: 35828
diff changeset
  1600
17149
e2b19c92ef51 Lemmas on dvd, power and finite summation added or strengthened.
ballarin
parents: 16733
diff changeset
  1601
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1602
subsection {* The formula for arithmetic sums *}
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1603
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1604
lemma gauss_sum:
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 55719
diff changeset
  1605
  "(2::'a::comm_semiring_1)*(\<Sum>i\<in>{1..n}. of_nat i) = of_nat n*((of_nat n)+1)"
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1606
proof (induct n)
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1607
  case 0
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1608
  show ?case by simp
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1609
next
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1610
  case (Suc n)
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1611
  then show ?case
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1612
    by (simp add: algebra_simps add: one_add_one [symmetric] del: one_add_one)
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1613
      (* FIXME: make numeral cancellation simprocs work for semirings *)
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1614
qed
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1615
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1616
theorem arith_series_general:
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1617
  "(2::'a::comm_semiring_1) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1618
  of_nat n * (a + (a + of_nat(n - 1)*d))"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1619
proof cases
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1620
  assume ngt1: "n > 1"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1621
  let ?I = "\<lambda>i. of_nat i" and ?n = "of_nat n"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1622
  have
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1623
    "(\<Sum>i\<in>{..<n}. a+?I i*d) =
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1624
     ((\<Sum>i\<in>{..<n}. a) + (\<Sum>i\<in>{..<n}. ?I i*d))"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1625
    by (rule setsum_addf)
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1626
  also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1627
  also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1628
    unfolding One_nat_def
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1629
    by (simp add: setsum_right_distrib atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt mult_ac)
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1630
  also have "2*\<dots> = 2*?n*a + d*2*(\<Sum>i\<in>{1..<n}. ?I i)"
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1631
    by (simp add: algebra_simps)
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1632
  also from ngt1 have "{1..<n} = {1..n - 1}"
28068
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1633
    by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost)
f6b2d1995171 cleaned up code generation for {.._} and {..<_}
nipkow
parents: 27656
diff changeset
  1634
  also from ngt1
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1635
  have "2*?n*a + d*2*(\<Sum>i\<in>{1..n - 1}. ?I i) = (2*?n*a + d*?I (n - 1)*?I n)"
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1636
    by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def)
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23413
diff changeset
  1637
       (simp add:  mult_ac trans [OF add_commute of_nat_Suc [symmetric]])
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1638
  finally show ?thesis
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1639
    unfolding mult_2 by (simp add: algebra_simps)
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1640
next
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1641
  assume "\<not>(n > 1)"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1642
  hence "n = 1 \<or> n = 0" by auto
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1643
  thus ?thesis by (auto simp: mult_2)
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1644
qed
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1645
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1646
lemma arith_series_nat:
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1647
  "(2::nat) * (\<Sum>i\<in>{..<n}. a+i*d) = n * (a + (a+(n - 1)*d))"
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1648
proof -
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1649
  have
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1650
    "2 * (\<Sum>i\<in>{..<n::nat}. a + of_nat(i)*d) =
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1651
    of_nat(n) * (a + (a + of_nat(n - 1)*d))"
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1652
    by (rule arith_series_general)
30079
293b896b9c25 make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
huffman
parents: 29960
diff changeset
  1653
  thus ?thesis
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35171
diff changeset
  1654
    unfolding One_nat_def by auto
19469
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1655
qed
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1656
958d2f2dd8d4 moved arithmetic series to geometric series in SetInterval
kleing
parents: 19376
diff changeset
  1657
lemma arith_series_int:
47222
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1658
  "2 * (\<Sum>i\<in>{..<n}. a + int i * d) = int n * (a + (a + int(n - 1)*d))"
1b7c909a6fad rephrase lemmas about arithmetic series using numeral '2'
huffman
parents: 47108
diff changeset
  1659
  by (fact arith_series_general) (* FIXME: duplicate *)
15418
e28853da5df5 removed two looping simplifications in SetInterval.thy; deleted the .ML file
paulson
parents: 15402
diff changeset
  1660
19022
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1661
lemma sum_diff_distrib:
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1662
  fixes P::"nat\<Rightarrow>nat"
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1663
  shows
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1664
  "\<forall>x. Q x \<le> P x  \<Longrightarrow>
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1665
  (\<Sum>x<n. P x) - (\<Sum>x<n. Q x) = (\<Sum>x<n. P x - Q x)"
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1666
proof (induct n)
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1667
  case 0 show ?case by simp
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1668
next
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1669
  case (Suc n)
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1670
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1671
  let ?lhs = "(\<Sum>x<n. P x) - (\<Sum>x<n. Q x)"
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1672
  let ?rhs = "\<Sum>x<n. P x - Q x"
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1673
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1674
  from Suc have "?lhs = ?rhs" by simp
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1675
  moreover
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1676
  from Suc have "?lhs + P n - Q n = ?rhs + (P n - Q n)" by simp
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1677
  moreover
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1678
  from Suc have
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1679
    "(\<Sum>x<n. P x) + P n - ((\<Sum>x<n. Q x) + Q n) = ?rhs + (P n - Q n)"
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1680
    by (subst diff_diff_left[symmetric],
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1681
        subst diff_add_assoc2)
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1682
       (auto simp: diff_add_assoc2 intro: setsum_mono)
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1683
  ultimately
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1684
  show ?case by simp
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1685
qed
0e6ec4fd204c * moved ThreeDivides from Isar_examples to better suited HOL/ex
kleing
parents: 17719
diff changeset
  1686
55718
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1687
lemma nat_diff_setsum_reindex:
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1688
  fixes x :: "'a::{comm_ring,monoid_mult}"
56193
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 55719
diff changeset
  1689
  shows "(\<Sum>i<n. f (n - Suc i)) = (\<Sum>i<n. f i)"
c726ecfb22b6 cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents: 55719
diff changeset
  1690
apply (subst setsum_reindex_cong [of "%i. n - Suc i" "{..< n}"])
55718
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1691
apply (auto simp: inj_on_def)
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1692
apply (rule_tac x="n - Suc x" in image_eqI, auto)
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1693
done
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1694
29960
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1695
subsection {* Products indexed over intervals *}
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1696
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1697
syntax
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1698
  "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1699
  "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1700
  "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _<_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1701
  "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(PROD _<=_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1702
syntax (xsymbols)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1703
  "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _.._./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1704
  "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _..<_./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1705
  "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_<_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1706
  "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_\<le>_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1707
syntax (HTML output)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1708
  "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _.._./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1709
  "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_ = _..<_./ _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1710
  "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_<_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1711
  "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Prod>_\<le>_./ _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1712
syntax (latex_prod output)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1713
  "_from_to_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1714
 ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{>_\<^raw:}$> _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1715
  "_from_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1716
 ("(3\<^raw:$\prod_{>_ = _\<^raw:}^{<>_\<^raw:}$> _)" [0,0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1717
  "_upt_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1718
 ("(3\<^raw:$\prod_{>_ < _\<^raw:}$> _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1719
  "_upto_setprod" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1720
 ("(3\<^raw:$\prod_{>_ \<le> _\<^raw:}$> _)" [0,0,10] 10)
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1721
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1722
translations
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1723
  "\<Prod>x=a..b. t" == "CONST setprod (%x. t) {a..b}"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1724
  "\<Prod>x=a..<b. t" == "CONST setprod (%x. t) {a..<b}"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1725
  "\<Prod>i\<le>n. t" == "CONST setprod (\<lambda>i. t) {..n}"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1726
  "\<Prod>i<n. t" == "CONST setprod (\<lambda>i. t) {..<n}"
9d5c6f376768 Syntactic support for products over set intervals
paulson
parents: 29920
diff changeset
  1727
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1728
subsection {* Transfer setup *}
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1729
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1730
lemma transfer_nat_int_set_functions:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1731
    "{..n} = nat ` {0..int n}"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1732
    "{m..n} = nat ` {int m..int n}"  (* need all variants of these! *)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1733
  apply (auto simp add: image_def)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1734
  apply (rule_tac x = "int x" in bexI)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1735
  apply auto
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1736
  apply (rule_tac x = "int x" in bexI)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1737
  apply auto
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1738
  done
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1739
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1740
lemma transfer_nat_int_set_function_closures:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1741
    "x >= 0 \<Longrightarrow> nat_set {x..y}"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1742
  by (simp add: nat_set_def)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1743
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35580
diff changeset
  1744
declare transfer_morphism_nat_int[transfer add
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1745
  return: transfer_nat_int_set_functions
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1746
    transfer_nat_int_set_function_closures
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1747
]
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1748
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1749
lemma transfer_int_nat_set_functions:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1750
    "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1751
  by (simp only: is_nat_def transfer_nat_int_set_functions
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1752
    transfer_nat_int_set_function_closures
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1753
    transfer_nat_int_set_return_embed nat_0_le
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1754
    cong: transfer_nat_int_set_cong)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1755
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1756
lemma transfer_int_nat_set_function_closures:
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1757
    "is_nat x \<Longrightarrow> nat_set {x..y}"
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1758
  by (simp only: transfer_nat_int_set_function_closures is_nat_def)
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1759
35644
d20cf282342e transfer: avoid camel case
haftmann
parents: 35580
diff changeset
  1760
declare transfer_morphism_int_nat[transfer add
33318
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1761
  return: transfer_int_nat_set_functions
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1762
    transfer_int_nat_set_function_closures
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1763
]
ddd97d9dfbfb moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
haftmann
parents: 33044
diff changeset
  1764
55242
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1765
lemma setprod_int_plus_eq: "setprod int {i..i+j} =  \<Prod>{int i..int (i+j)}"
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1766
  by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv)
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1767
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1768
lemma setprod_int_eq: "setprod int {i..j} =  \<Prod>{int i..int j}"
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1769
proof (cases "i \<le> j")
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1770
  case True
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1771
  then show ?thesis
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1772
    by (metis Nat.le_iff_add setprod_int_plus_eq)
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1773
next
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1774
  case False
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1775
  then show ?thesis
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1776
    by auto
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1777
qed
413ec965f95d Number_Theory: prime is no longer overloaded, but only for nat. Automatic coercion to int enabled.
paulson <lp15@cam.ac.uk>
parents: 55143
diff changeset
  1778
55718
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1779
lemma setprod_power_distrib:
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1780
  fixes f :: "'a \<Rightarrow> 'b::comm_semiring_1"
55719
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1781
  shows "setprod f A ^ n = setprod (\<lambda>x. (f x)^n) A"
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1782
proof (cases "finite A") 
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1783
  case True then show ?thesis 
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1784
    by (induct A rule: finite_induct) (auto simp add: power_mult_distrib)
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1785
next
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1786
  case False then show ?thesis 
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1787
    by (metis setprod_infinite power_one)
cdddd073bff8 Lemmas about Reals, norm, etc., and cleaner variants of existing ones
paulson <lp15@cam.ac.uk>
parents: 55718
diff changeset
  1788
qed
55718
34618f031ba9 A few lemmas about summations, etc.
paulson <lp15@cam.ac.uk>
parents: 55242
diff changeset
  1789
8924
c434283b4cfa Added SetInterval
nipkow
parents:
diff changeset
  1790
end