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