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