src/HOL/Lattices_Big.thy
author wenzelm
Thu, 25 Sep 2025 22:26:31 +0200
changeset 83231 06a05e098347
parent 82858 52cf8f3f1652
permissions -rw-r--r--
more Java heap space for the sake of Isabelle/AFP cluster build stability, but this is better changed elsewhere; enforce rebuild of Isabelle/ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Lattices_Big.thy
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     2
    Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     3
                with contributions by Jeremy Avigad
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     4
*)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     5
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
     6
section \<open>Big infimum (minimum) and supremum (maximum) over finite (non-empty) sets\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     7
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
     8
theory Lattices_Big
74979
4d77dd3019d1 earlier availability of lifting
haftmann
parents: 73411
diff changeset
     9
  imports Groups_Big Option
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    10
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    11
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
    12
subsection \<open>Generic lattice operations over a set\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    13
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
    14
subsubsection \<open>Without neutral element\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    15
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    16
locale semilattice_set = semilattice
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    17
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    18
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    19
interpretation comp_fun_idem f
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60758
diff changeset
    20
  by standard (simp_all add: fun_eq_iff left_commute)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    21
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    22
definition F :: "'a set \<Rightarrow> 'a"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    23
where
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    24
  eq_fold': "F A = the (Finite_Set.fold (\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)) None A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    25
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    26
lemma eq_fold:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    27
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    28
  shows "F (insert x A) = Finite_Set.fold f x A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    29
proof (rule sym)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    30
  let ?f = "\<lambda>x y. Some (case y of None \<Rightarrow> x | Some z \<Rightarrow> f x z)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    31
  interpret comp_fun_idem "?f"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60758
diff changeset
    32
    by standard (simp_all add: fun_eq_iff commute left_commute split: option.split)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    33
  from assms show "Finite_Set.fold f x A = F (insert x A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    34
  proof induct
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    35
    case empty then show ?case by (simp add: eq_fold')
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    36
  next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    37
    case (insert y B) then show ?case by (simp add: insert_commute [of x] eq_fold')
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    38
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    39
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    40
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    41
lemma singleton [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    42
  "F {x} = x"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    43
  by (simp add: eq_fold)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    44
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    45
lemma insert_not_elem:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    46
  assumes "finite A" and "x \<notin> A" and "A \<noteq> {}"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    47
  shows "F (insert x A) = x \<^bold>* F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    48
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
    49
  from \<open>A \<noteq> {}\<close> obtain b where "b \<in> A" by blast
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    50
  then obtain B where *: "A = insert b B" "b \<notin> B" by (blast dest: mk_disjoint_insert)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
    51
  with \<open>finite A\<close> and \<open>x \<notin> A\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    52
    have "finite (insert x B)" and "b \<notin> insert x B" by auto
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    53
  then have "F (insert b (insert x B)) = x \<^bold>* F (insert b B)"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    54
    by (simp add: eq_fold)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    55
  then show ?thesis by (simp add: * insert_commute)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    56
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    57
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    58
lemma in_idem:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    59
  assumes "finite A" and "x \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    60
  shows "x \<^bold>* F A = F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    61
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    62
  from assms have "A \<noteq> {}" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
    63
  with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    64
    by (induct A rule: finite_ne_induct) (auto simp add: ac_simps insert_not_elem)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    65
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    66
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    67
lemma insert [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    68
  assumes "finite A" and "A \<noteq> {}"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    69
  shows "F (insert x A) = x \<^bold>* F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    70
  using assms by (cases "x \<in> A") (simp_all add: insert_absorb in_idem insert_not_elem)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    71
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    72
lemma union:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    73
  assumes "finite A" "A \<noteq> {}" and "finite B" "B \<noteq> {}"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    74
  shows "F (A \<union> B) = F A \<^bold>* F B"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    75
  using assms by (induct A rule: finite_ne_induct) (simp_all add: ac_simps)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    76
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    77
lemma remove:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    78
  assumes "finite A" and "x \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    79
  shows "F A = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    80
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    81
  from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    82
  with assms show ?thesis by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    83
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    84
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    85
lemma insert_remove:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    86
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    87
  shows "F (insert x A) = (if A - {x} = {} then x else x \<^bold>* F (A - {x}))"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    88
  using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    89
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    90
lemma subset:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    91
  assumes "finite A" "B \<noteq> {}" and "B \<subseteq> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    92
  shows "F B \<^bold>* F A = F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    93
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    94
  from assms have "A \<noteq> {}" and "finite B" by (auto dest: finite_subset)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    95
  with assms show ?thesis by (simp add: union [symmetric] Un_absorb1)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    96
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    97
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
    98
lemma closed:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
    99
  assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x \<^bold>* y \<in> {x, y}"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   100
  shows "F A \<in> A"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   101
using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   102
  case singleton then show ?case by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   103
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   104
  case insert with elem show ?case by force
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   105
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   106
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   107
lemma hom_commute:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   108
  assumes hom: "\<And>x y. h (x \<^bold>* y) = h x \<^bold>* h y"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   109
  and N: "finite N" "N \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   110
  shows "h (F N) = F (h ` N)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   111
using N proof (induct rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   112
  case singleton thus ?case by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   113
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   114
  case (insert n N)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   115
  then have "h (F (insert n N)) = h (n \<^bold>* F N)" by simp
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   116
  also have "\<dots> = h n \<^bold>* h (F N)" by (rule hom)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   117
  also have "h (F N) = F (h ` N)" by (rule insert)
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   118
  also have "h n \<^bold>* \<dots> = F (insert (h n) (h ` N))"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   119
    using insert by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   120
  also have "insert (h n) (h ` N) = h ` insert n N" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   121
  finally show ?case .
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   122
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   123
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56140
diff changeset
   124
lemma infinite: "\<not> finite A \<Longrightarrow> F A = the None"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56140
diff changeset
   125
  unfolding eq_fold' by (cases "finite (UNIV::'a set)") (auto intro: finite_subset fold_infinite)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56140
diff changeset
   126
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   127
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   128
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   129
locale semilattice_order_set = binary?: semilattice_order + semilattice_set
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   130
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   131
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   132
lemma bounded_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   133
  assumes "finite A" and "A \<noteq> {}"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   134
  shows "x \<^bold>\<le> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<^bold>\<le> a)"
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   135
  using assms by (induct rule: finite_ne_induct) simp_all
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   136
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   137
lemma boundedI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   138
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   139
  assumes "A \<noteq> {}"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   140
  assumes "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   141
  shows "x \<^bold>\<le> F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   142
  using assms by (simp add: bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   143
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   144
lemma boundedE:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   145
  assumes "finite A" and "A \<noteq> {}" and "x \<^bold>\<le> F A"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   146
  obtains "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   147
  using assms by (simp add: bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   148
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   149
lemma coboundedI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   150
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   151
    and "a \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   152
  shows "F A \<^bold>\<le> a"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   153
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   154
  from assms have "A \<noteq> {}" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   155
  from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   156
  proof (induct rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   157
    case singleton thus ?case by (simp add: refl)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   158
  next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   159
    case (insert x B)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   160
    from insert have "a = x \<or> a \<in> B" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   161
    then show ?case using insert by (auto intro: coboundedI2)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   162
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   163
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   164
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   165
lemma subset_imp:
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   166
  assumes "A \<subseteq> B" and "A \<noteq> {}" and "finite B"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   167
  shows "F B \<^bold>\<le> F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   168
proof (cases "A = B")
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   169
  case True then show ?thesis by (simp add: refl)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   170
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   171
  case False
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   172
  have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   173
  then have "F B = F (A \<union> (B - A))" by simp
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   174
  also have "\<dots> = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset)
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   175
  also have "\<dots> \<^bold>\<le> F A" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   176
  finally show ?thesis .
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   177
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   178
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   179
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   180
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   181
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   182
subsubsection \<open>With neutral element\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   183
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   184
locale semilattice_neutr_set = semilattice_neutr
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   185
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   186
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   187
interpretation comp_fun_idem f
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60758
diff changeset
   188
  by standard (simp_all add: fun_eq_iff left_commute)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   189
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   190
definition F :: "'a set \<Rightarrow> 'a"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   191
where
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   192
  eq_fold: "F A = Finite_Set.fold f \<^bold>1 A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   193
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   194
lemma infinite [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   195
  "\<not> finite A \<Longrightarrow> F A = \<^bold>1"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   196
  by (simp add: eq_fold)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   197
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   198
lemma empty [simp]:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   199
  "F {} = \<^bold>1"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   200
  by (simp add: eq_fold)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   201
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   202
lemma insert [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   203
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   204
  shows "F (insert x A) = x \<^bold>* F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   205
  using assms by (simp add: eq_fold)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   206
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   207
lemma in_idem:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   208
  assumes "finite A" and "x \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   209
  shows "x \<^bold>* F A = F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   210
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   211
  from assms have "A \<noteq> {}" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   212
  with \<open>finite A\<close> show ?thesis using \<open>x \<in> A\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   213
    by (induct A rule: finite_ne_induct) (auto simp add: ac_simps)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   214
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   215
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   216
lemma union:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   217
  assumes "finite A" and "finite B"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   218
  shows "F (A \<union> B) = F A \<^bold>* F B"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   219
  using assms by (induct A) (simp_all add: ac_simps)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   220
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   221
lemma remove:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   222
  assumes "finite A" and "x \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   223
  shows "F A = x \<^bold>* F (A - {x})"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   224
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   225
  from assms obtain B where "A = insert x B" and "x \<notin> B" by (blast dest: mk_disjoint_insert)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   226
  with assms show ?thesis by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   227
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   228
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   229
lemma insert_remove:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   230
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   231
  shows "F (insert x A) = x \<^bold>* F (A - {x})"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   232
  using assms by (cases "x \<in> A") (simp_all add: insert_absorb remove)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   233
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   234
lemma subset:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   235
  assumes "finite A" and "B \<subseteq> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   236
  shows "F B \<^bold>* F A = F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   237
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   238
  from assms have "finite B" by (auto dest: finite_subset)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   239
  with assms show ?thesis by (simp add: union [symmetric] Un_absorb1)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   240
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   241
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   242
lemma closed:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   243
  assumes "finite A" "A \<noteq> {}" and elem: "\<And>x y. x \<^bold>* y \<in> {x, y}"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   244
  shows "F A \<in> A"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   245
using \<open>finite A\<close> \<open>A \<noteq> {}\<close> proof (induct rule: finite_ne_induct)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   246
  case singleton then show ?case by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   247
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   248
  case insert with elem show ?case by force
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   249
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   250
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   251
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   252
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   253
locale semilattice_order_neutr_set = binary?: semilattice_neutr_order + semilattice_neutr_set
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   254
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   255
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   256
lemma bounded_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   257
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   258
  shows "x \<^bold>\<le> F A \<longleftrightarrow> (\<forall>a\<in>A. x \<^bold>\<le> a)"
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   259
  using assms by (induct A) simp_all
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   260
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   261
lemma boundedI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   262
  assumes "finite A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   263
  assumes "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   264
  shows "x \<^bold>\<le> F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   265
  using assms by (simp add: bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   266
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   267
lemma boundedE:
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   268
  assumes "finite A" and "x \<^bold>\<le> F A"
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   269
  obtains "\<And>a. a \<in> A \<Longrightarrow> x \<^bold>\<le> a"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   270
  using assms by (simp add: bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   271
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   272
lemma coboundedI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   273
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   274
    and "a \<in> A"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   275
  shows "F A \<^bold>\<le> a"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   276
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   277
  from assms have "A \<noteq> {}" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   278
  from \<open>finite A\<close> \<open>A \<noteq> {}\<close> \<open>a \<in> A\<close> show ?thesis
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   279
  proof (induct rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   280
    case singleton thus ?case by (simp add: refl)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   281
  next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   282
    case (insert x B)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   283
    from insert have "a = x \<or> a \<in> B" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   284
    then show ?case using insert by (auto intro: coboundedI2)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   285
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   286
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   287
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   288
lemma subset_imp:
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   289
  assumes "A \<subseteq> B" and "finite B"
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   290
  shows "F B \<^bold>\<le> F A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   291
proof (cases "A = B")
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   292
  case True then show ?thesis by (simp add: refl)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   293
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   294
  case False
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   295
  have B: "B = A \<union> (B - A)" using \<open>A \<subseteq> B\<close> by blast
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   296
  then have "F B = F (A \<union> (B - A))" by simp
63290
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   297
  also have "\<dots> = F A \<^bold>* F (B - A)" using False assms by (subst union) (auto intro: finite_subset)
9ac558ab0906 boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents: 61776
diff changeset
   298
  also have "\<dots> \<^bold>\<le> F A" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   299
  finally show ?thesis .
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   300
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   301
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   302
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   303
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   304
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   305
subsection \<open>Lattice operations on finite sets\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   306
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   307
context semilattice_inf
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   308
begin
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   309
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   310
sublocale Inf_fin: semilattice_order_set inf less_eq less
61776
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
   311
defines
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 80770
diff changeset
   312
  Inf_fin (\<open>\<Sqinter>\<^sub>f\<^sub>i\<^sub>n _\<close> [900] 900) = Inf_fin.F ..
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   313
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   314
end
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   315
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   316
context semilattice_sup
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   317
begin
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   318
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   319
sublocale Sup_fin: semilattice_order_set sup greater_eq greater
61776
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
   320
defines
80932
261cd8722677 standardize mixfix annotations via "isabelle update -u mixfix_cartouches -l Pure HOL" --- to simplify systematic editing;
wenzelm
parents: 80770
diff changeset
   321
  Sup_fin (\<open>\<Squnion>\<^sub>f\<^sub>i\<^sub>n _\<close> [900] 900) = Sup_fin.F ..
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   322
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   323
end
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   324
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   325
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   326
subsection \<open>Infimum and Supremum over non-empty sets\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   327
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   328
context lattice
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   329
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   330
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   331
lemma Inf_fin_le_Sup_fin [simp]: 
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   332
  assumes "finite A" and "A \<noteq> {}"
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   333
  shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA"
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   334
proof -
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   335
  from \<open>A \<noteq> {}\<close> obtain a where "a \<in> A" by blast
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   336
  with \<open>finite A\<close> have "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<le> a" by (rule Inf_fin.coboundedI)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   337
  moreover from \<open>finite A\<close> \<open>a \<in> A\<close> have "a \<le> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA" by (rule Sup_fin.coboundedI)
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   338
  ultimately show ?thesis by (rule order_trans)
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   339
qed
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   340
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   341
lemma sup_Inf_absorb [simp]:
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   342
  "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> \<Sqinter>\<^sub>f\<^sub>i\<^sub>nA \<squnion> a = a"
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   343
  by (rule sup_absorb2) (rule Inf_fin.coboundedI)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   344
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   345
lemma inf_Sup_absorb [simp]:
54745
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   346
  "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> a \<sqinter> \<Squnion>\<^sub>f\<^sub>i\<^sub>nA = a"
46e441e61ff5 disambiguation of interpretation prefixes
haftmann
parents: 54744
diff changeset
   347
  by (rule inf_absorb1) (rule Sup_fin.coboundedI)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   348
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   349
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   350
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   351
context distrib_lattice
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   352
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   353
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   354
lemma sup_Inf1_distrib:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   355
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   356
    and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   357
  shows "sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x a|a. a \<in> A}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   358
using assms by (simp add: image_def Inf_fin.hom_commute [where h="sup x", OF sup_inf_distrib1])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   359
  (rule arg_cong [where f="Inf_fin"], blast)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   360
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   361
lemma sup_Inf2_distrib:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   362
  assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   363
  shows "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   364
using A proof (induct rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   365
  case singleton then show ?case
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   366
    by (simp add: sup_Inf1_distrib [OF B])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   367
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   368
  case (insert x A)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   369
  have finB: "finite {sup x b |b. b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   370
    by (rule finite_surj [where f = "sup x", OF B(1)], auto)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   371
  have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   372
  proof -
69276
3d954183b707 replaced some ancient ASCII syntax
haftmann
parents: 68980
diff changeset
   373
    have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (\<Union>a\<in>A. \<Union>b\<in>B. {sup a b})"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   374
      by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   375
    thus ?thesis by(simp add: insert(1) B(1))
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   376
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   377
  have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   378
  have "sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB) = sup (inf x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA)) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   379
    using insert by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   380
  also have "\<dots> = inf (sup x (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB)) (sup (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>nB))" by(rule sup_inf_distrib2)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   381
  also have "\<dots> = inf (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup x b|b. b \<in> B}) (\<Sqinter>\<^sub>f\<^sub>i\<^sub>n{sup a b|a b. a \<in> A \<and> b \<in> B})"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   382
    using insert by(simp add:sup_Inf1_distrib[OF B])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   383
  also have "\<dots> = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   384
    (is "_ = \<Sqinter>\<^sub>f\<^sub>i\<^sub>n?M")
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   385
    using B insert
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   386
    by (simp add: Inf_fin.union [OF finB _ finAB ne])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   387
  also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   388
    by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   389
  finally show ?case .
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   390
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   391
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   392
lemma inf_Sup1_distrib:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   393
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   394
  shows "inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x a|a. a \<in> A}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   395
using assms by (simp add: image_def Sup_fin.hom_commute [where h="inf x", OF inf_sup_distrib1])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   396
  (rule arg_cong [where f="Sup_fin"], blast)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   397
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   398
lemma inf_Sup2_distrib:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   399
  assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   400
  shows "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = \<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   401
using A proof (induct rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   402
  case singleton thus ?case
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   403
    by(simp add: inf_Sup1_distrib [OF B])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   404
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   405
  case (insert x A)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   406
  have finB: "finite {inf x b |b. b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   407
    by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   408
  have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   409
  proof -
69276
3d954183b707 replaced some ancient ASCII syntax
haftmann
parents: 68980
diff changeset
   410
    have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (\<Union>a\<in>A. \<Union>b\<in>B. {inf a b})"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   411
      by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   412
    thus ?thesis by(simp add: insert(1) B(1))
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   413
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   414
  have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   415
  have "inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>n(insert x A)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB) = inf (sup x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA)) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   416
    using insert by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   417
  also have "\<dots> = sup (inf x (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB)) (inf (\<Squnion>\<^sub>f\<^sub>i\<^sub>nA) (\<Squnion>\<^sub>f\<^sub>i\<^sub>nB))" by(rule inf_sup_distrib2)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   418
  also have "\<dots> = sup (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf x b|b. b \<in> B}) (\<Squnion>\<^sub>f\<^sub>i\<^sub>n{inf a b|a b. a \<in> A \<and> b \<in> B})"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   419
    using insert by(simp add:inf_Sup1_distrib[OF B])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   420
  also have "\<dots> = \<Squnion>\<^sub>f\<^sub>i\<^sub>n({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   421
    (is "_ = \<Squnion>\<^sub>f\<^sub>i\<^sub>n?M")
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   422
    using B insert
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   423
    by (simp add: Sup_fin.union [OF finB _ finAB ne])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   424
  also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   425
    by blast
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   426
  finally show ?case .
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   427
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   428
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   429
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   430
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   431
context complete_lattice
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   432
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   433
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   434
lemma Inf_fin_Inf:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   435
  assumes "finite A" and "A \<noteq> {}"
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   436
  shows "\<Sqinter>\<^sub>f\<^sub>i\<^sub>nA = \<Sqinter>A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   437
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   438
  from assms obtain b B where "A = insert b B" and "finite B" by auto
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   439
  then show ?thesis
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   440
    by (simp add: Inf_fin.eq_fold inf_Inf_fold_inf inf.commute [of b])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   441
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   442
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   443
lemma Sup_fin_Sup:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   444
  assumes "finite A" and "A \<noteq> {}"
54868
bab6cade3cc5 prefer target-style syntaxx for sublocale
haftmann
parents: 54864
diff changeset
   445
  shows "\<Squnion>\<^sub>f\<^sub>i\<^sub>nA = \<Squnion>A"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   446
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   447
  from assms obtain b B where "A = insert b B" and "finite B" by auto
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   448
  then show ?thesis
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   449
    by (simp add: Sup_fin.eq_fold sup_Sup_fold_sup sup.commute [of b])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   450
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   451
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   452
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   453
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   454
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   455
subsection \<open>Minimum and Maximum over non-empty sets\<close>
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   456
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   457
context linorder
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   458
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   459
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   460
sublocale Min: semilattice_order_set min less_eq less
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   461
  + Max: semilattice_order_set max greater_eq greater
61776
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
   462
defines
57bb7da5c867 modernized
haftmann
parents: 61605
diff changeset
   463
  Min = Min.F and Max = Max.F ..
54864
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   464
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   465
end
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   466
67969
83c8cafdebe8 Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   467
syntax
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   468
  "_MIN1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           (\<open>(\<open>indent=3 notation=\<open>binder MIN\<close>\<close>MIN _./ _)\<close> [0, 10] 10)
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   469
  "_MIN"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  (\<open>(\<open>indent=3 notation=\<open>binder MIN\<close>\<close>MIN _\<in>_./ _)\<close> [0, 0, 10] 10)
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   470
  "_MAX1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           (\<open>(\<open>indent=3 notation=\<open>binder MAX\<close>\<close>MAX _./ _)\<close> [0, 10] 10)
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   471
  "_MAX"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  (\<open>(\<open>indent=3 notation=\<open>binder MAX\<close>\<close>MAX _\<in>_./ _)\<close> [0, 0, 10] 10)
67969
83c8cafdebe8 Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   472
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   473
syntax_consts
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   474
  "_MIN1" "_MIN" \<rightleftharpoons> Min and
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   475
  "_MAX1" "_MAX" \<rightleftharpoons> Max
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   476
67969
83c8cafdebe8 Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   477
translations
68796
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   478
  "MIN x y. f"   \<rightleftharpoons> "MIN x. MIN y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   479
  "MIN x. f"     \<rightleftharpoons> "CONST Min (CONST range (\<lambda>x. f))"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   480
  "MIN x\<in>A. f"   \<rightleftharpoons> "CONST Min ((\<lambda>x. f) ` A)"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   481
  "MAX x y. f"   \<rightleftharpoons> "MAX x. MAX y. f"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   482
  "MAX x. f"     \<rightleftharpoons> "CONST Max (CONST range (\<lambda>x. f))"
9ca183045102 simplified syntax setup for big operators under image, retaining input abbreviations for backward compatibility
haftmann
parents: 68795
diff changeset
   483
  "MAX x\<in>A. f"   \<rightleftharpoons> "CONST Max ((\<lambda>x. f) ` A)"
67969
83c8cafdebe8 Syntax for the special cases Min(A`I) and Max (A`I)
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   484
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   485
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69276
diff changeset
   486
text \<open>An aside: \<^const>\<open>Min\<close>/\<^const>\<open>Max\<close> on linear orders as special case of \<^const>\<open>Inf_fin\<close>/\<^const>\<open>Sup_fin\<close>\<close>
54864
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   487
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   488
lemma Inf_fin_Min:
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   489
  "Inf_fin = (Min :: 'a::{semilattice_inf, linorder} set \<Rightarrow> 'a)"
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   490
  by (simp add: Inf_fin_def Min_def inf_min)
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   491
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   492
lemma Sup_fin_Max:
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   493
  "Sup_fin = (Max :: 'a::{semilattice_sup, linorder} set \<Rightarrow> 'a)"
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   494
  by (simp add: Sup_fin_def Max_def sup_max)
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   495
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   496
context linorder
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   497
begin
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54863
diff changeset
   498
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   499
lemma dual_min:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   500
  "ord.min greater_eq = max"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   501
  by (auto simp add: ord.min_def max_def fun_eq_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   502
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   503
lemma dual_max:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   504
  "ord.max greater_eq = min"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   505
  by (auto simp add: ord.max_def min_def fun_eq_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   506
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   507
lemma dual_Min:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   508
  "linorder.Min greater_eq = Max"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   509
proof -
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   510
  interpret dual: linorder greater_eq greater by (fact dual_linorder)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   511
  show ?thesis by (simp add: dual.Min_def dual_min Max_def)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   512
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   513
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   514
lemma dual_Max:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   515
  "linorder.Max greater_eq = Min"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   516
proof -
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61566
diff changeset
   517
  interpret dual: linorder greater_eq greater by (fact dual_linorder)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   518
  show ?thesis by (simp add: dual.Max_def dual_max Min_def)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   519
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   520
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   521
lemmas Min_singleton = Min.singleton
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   522
lemmas Max_singleton = Max.singleton
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   523
lemmas Min_insert = Min.insert
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   524
lemmas Max_insert = Max.insert
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   525
lemmas Min_Un = Min.union
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   526
lemmas Max_Un = Max.union
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   527
lemmas hom_Min_commute = Min.hom_commute
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   528
lemmas hom_Max_commute = Max.hom_commute
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   529
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   530
lemma Min_in [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   531
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   532
  shows "Min A \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   533
  using assms by (auto simp add: min_def Min.closed)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   534
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   535
lemma Max_in [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   536
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   537
  shows "Max A \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   538
  using assms by (auto simp add: max_def Max.closed)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   539
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   540
lemma Min_insert2:
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   541
  assumes "finite A" and min: "\<And>b. b \<in> A \<Longrightarrow> a \<le> b"
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   542
  shows "Min (insert a A) = a"
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   543
proof (cases "A = {}")
63915
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   544
  case True
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   545
  then show ?thesis by simp
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   546
next
63915
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   547
  case False
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   548
  with \<open>finite A\<close> have "Min (insert a A) = min a (Min A)"
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   549
    by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   550
  moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> min have "a \<le> Min A" by simp
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   551
  ultimately show ?thesis by (simp add: min.absorb1)
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   552
qed
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   553
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   554
lemma Max_insert2:
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   555
  assumes "finite A" and max: "\<And>b. b \<in> A \<Longrightarrow> b \<le> a"
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   556
  shows "Max (insert a A) = a"
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   557
proof (cases "A = {}")
63915
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   558
  case True
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   559
  then show ?thesis by simp
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   560
next
63915
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   561
  case False
bab633745c7f tuned proofs;
wenzelm
parents: 63290
diff changeset
   562
  with \<open>finite A\<close> have "Max (insert a A) = max a (Max A)"
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   563
    by simp
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   564
  moreover from \<open>finite A\<close> \<open>A \<noteq> {}\<close> max have "Max A \<le> a" by simp
58467
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   565
  ultimately show ?thesis by (simp add: max.absorb1)
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   566
qed
6a3da58f7233 moved to HOL and generalized
haftmann
parents: 57800
diff changeset
   567
73221
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   568
lemma Max_const[simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Max ((\<lambda>_. c) ` A) = c"
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   569
using Max_in image_is_empty by blast
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   570
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   571
lemma Min_const[simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Min ((\<lambda>_. c) ` A) = c"
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   572
using Min_in image_is_empty by blast
b1aa641eee4c added lemmas
nipkow
parents: 73102
diff changeset
   573
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   574
lemma Min_le [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   575
  assumes "finite A" and "x \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   576
  shows "Min A \<le> x"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   577
  using assms by (fact Min.coboundedI)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   578
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   579
lemma Max_ge [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   580
  assumes "finite A" and "x \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   581
  shows "x \<le> Max A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   582
  using assms by (fact Max.coboundedI)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   583
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   584
lemma Min_eqI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   585
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   586
  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   587
    and "x \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   588
  shows "Min A = x"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   589
proof (rule order.antisym)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   590
  from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   591
  with assms show "Min A \<ge> x" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   592
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   593
  from assms show "x \<ge> Min A" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   594
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   595
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   596
lemma Max_eqI:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   597
  assumes "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   598
  assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   599
    and "x \<in> A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   600
  shows "Max A = x"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   601
proof (rule order.antisym)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   602
  from \<open>x \<in> A\<close> have "A \<noteq> {}" by auto
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   603
  with assms show "Max A \<le> x" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   604
next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   605
  from assms show "x \<le> Max A" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   606
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   607
66425
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   608
lemma eq_Min_iff:
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   609
  "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> m = Min A  \<longleftrightarrow>  m \<in> A \<and> (\<forall>a \<in> A. m \<le> a)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   610
by (meson Min_in Min_le order.antisym)
66425
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   611
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   612
lemma Min_eq_iff:
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   613
  "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Min A = m  \<longleftrightarrow>  m \<in> A \<and> (\<forall>a \<in> A. m \<le> a)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   614
by (meson Min_in Min_le order.antisym)
66425
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   615
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   616
lemma eq_Max_iff:
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   617
  "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> m = Max A  \<longleftrightarrow>  m \<in> A \<and> (\<forall>a \<in> A. a \<le> m)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   618
by (meson Max_in Max_ge order.antisym)
66425
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   619
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   620
lemma Max_eq_iff:
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   621
  "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> Max A = m  \<longleftrightarrow>  m \<in> A \<and> (\<forall>a \<in> A. a \<le> m)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   622
by (meson Max_in Max_ge order.antisym)
66425
8756322dc5de added Min_mset and Max_mset
nipkow
parents: 66364
diff changeset
   623
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   624
context
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   625
  fixes A :: "'a set"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   626
  assumes fin_nonempty: "finite A" "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   627
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   628
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   629
lemma Min_ge_iff [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   630
  "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   631
  using fin_nonempty by (fact Min.bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   632
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   633
lemma Max_le_iff [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   634
  "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   635
  using fin_nonempty by (fact Max.bounded_iff)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   636
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   637
lemma Min_gr_iff [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   638
  "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   639
  using fin_nonempty  by (induct rule: finite_ne_induct) simp_all
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   640
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   641
lemma Max_less_iff [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   642
  "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   643
  using fin_nonempty by (induct rule: finite_ne_induct) simp_all
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   644
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   645
lemma Min_le_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   646
  "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   647
  using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   648
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   649
lemma Max_ge_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   650
  "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   651
  using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   652
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   653
lemma Min_less_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   654
  "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   655
  using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   656
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   657
lemma Max_gr_iff:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   658
  "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   659
  using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   660
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   661
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   662
80175
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   663
text \<open>Handy results about @{term Max} and @{term Min} by Chelsea Edmonds\<close>
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   664
lemma obtains_Max:
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   665
  assumes "finite A" and "A \<noteq> {}"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   666
  obtains x where "x \<in> A" and "Max A = x"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   667
  using assms Max_in by blast
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   668
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   669
lemma obtains_MAX:
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   670
  assumes "finite A" and "A \<noteq> {}"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   671
  obtains x where "x \<in> A" and "Max (f ` A) = f x"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   672
  using obtains_Max
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   673
  by (metis (mono_tags, opaque_lifting) assms(1) assms(2) empty_is_image finite_imageI image_iff) 
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   674
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   675
lemma obtains_Min:
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   676
  assumes "finite A" and "A \<noteq> {}"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   677
  obtains x where "x \<in> A" and "Min A = x"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   678
  using assms Min_in by blast
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   679
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   680
lemma obtains_MIN:
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   681
  assumes "finite A" and "A \<noteq> {}"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   682
  obtains x where "x \<in> A" and "Min (f ` A) = f x"
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   683
  using obtains_Min assms empty_is_image finite_imageI image_iff
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   684
  by (metis (mono_tags, opaque_lifting)) 
200107cdd3ac Some new simprules – and patches for proofs
paulson <lp15@cam.ac.uk>
parents: 75669
diff changeset
   685
57800
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   686
lemma Max_eq_if:
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   687
  assumes "finite A"  "finite B"  "\<forall>a\<in>A. \<exists>b\<in>B. a \<le> b"  "\<forall>b\<in>B. \<exists>a\<in>A. b \<le> a"
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   688
  shows "Max A = Max B"
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   689
proof cases
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   690
  assume "A = {}" thus ?thesis using assms by simp
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   691
next
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   692
  assume "A \<noteq> {}" thus ?thesis using assms
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   693
    by(blast intro: order.antisym Max_in Max_ge_iff[THEN iffD2])
57800
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   694
qed
84748234de9d added lemma
nipkow
parents: 56993
diff changeset
   695
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   696
lemma Min_antimono:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   697
  assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   698
  shows "Min N \<le> Min M"
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   699
  using assms by (fact Min.subset_imp)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   700
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   701
lemma Max_mono:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   702
  assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   703
  shows "Max M \<le> Max N"
67525
5d04d7bcd5f6 avoid concrete (anti)mono in theorem names since it could be the other way round
haftmann
parents: 67169
diff changeset
   704
  using assms by (fact Max.subset_imp)
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   705
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   706
lemma mono_Min_commute:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   707
  assumes "mono f"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   708
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   709
  shows "f (Min A) = Min (f ` A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   710
proof (rule linorder_class.Min_eqI [symmetric])
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   711
  from \<open>finite A\<close> show "finite (f ` A)" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   712
  from assms show "f (Min A) \<in> f ` A" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   713
  fix x
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   714
  assume "x \<in> f ` A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   715
  then obtain y where "y \<in> A" and "x = f y" ..
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   716
  with assms have "Min A \<le> y" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   717
  with \<open>mono f\<close> have "f (Min A) \<le> f y" by (rule monoE)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   718
  with \<open>x = f y\<close> show "f (Min A) \<le> x" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   719
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   720
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   721
lemma mono_Max_commute:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   722
  assumes "mono f"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   723
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   724
  shows "f (Max A) = Max (f ` A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   725
proof (rule linorder_class.Max_eqI [symmetric])
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   726
  from \<open>finite A\<close> show "finite (f ` A)" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   727
  from assms show "f (Max A) \<in> f ` A" by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   728
  fix x
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   729
  assume "x \<in> f ` A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   730
  then obtain y where "y \<in> A" and "x = f y" ..
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   731
  with assms have "y \<le> Max A" by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   732
  with \<open>mono f\<close> have "f y \<le> f (Max A)" by (rule monoE)
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   733
  with \<open>x = f y\<close> show "x \<le> f (Max A)" by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   734
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   735
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   736
lemma finite_linorder_max_induct [consumes 1, case_names empty insert]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   737
  assumes fin: "finite A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   738
  and empty: "P {}" 
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   739
  and insert: "\<And>b A. finite A \<Longrightarrow> \<forall>a\<in>A. a < b \<Longrightarrow> P A \<Longrightarrow> P (insert b A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   740
  shows "P A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   741
using fin empty insert
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   742
proof (induct rule: finite_psubset_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   743
  case (psubset A)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   744
  have IH: "\<And>B. \<lbrakk>B < A; P {}; (\<And>A b. \<lbrakk>finite A; \<forall>a\<in>A. a<b; P A\<rbrakk> \<Longrightarrow> P (insert b A))\<rbrakk> \<Longrightarrow> P B" by fact 
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   745
  have fin: "finite A" by fact 
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   746
  have empty: "P {}" by fact
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   747
  have step: "\<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)" by fact
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   748
  show "P A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   749
  proof (cases "A = {}")
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   750
    assume "A = {}" 
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   751
    then show "P A" using \<open>P {}\<close> by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   752
  next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   753
    let ?B = "A - {Max A}" 
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   754
    let ?A = "insert (Max A) ?B"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   755
    have "finite ?B" using \<open>finite A\<close> by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   756
    assume "A \<noteq> {}"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67566
diff changeset
   757
    with \<open>finite A\<close> have "Max A \<in> A" by auto
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   758
    then have A: "?A = A" using insert_Diff_single insert_absorb by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   759
    then have "P ?B" using \<open>P {}\<close> step IH [of ?B] by blast
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   760
    moreover 
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   761
    have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF \<open>finite A\<close>] by fastforce
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   762
    ultimately show "P A" using A insert_Diff_single step [OF \<open>finite ?B\<close>] by fastforce
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   763
  qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   764
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   765
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   766
lemma finite_linorder_min_induct [consumes 1, case_names empty insert]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   767
  "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   768
  by (rule linorder.finite_linorder_max_induct [OF dual_linorder])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   769
70184
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   770
lemma finite_ranking_induct[consumes 1, case_names empty insert]:
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   771
  fixes f :: "'b \<Rightarrow> 'a"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   772
  assumes "finite S"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   773
  assumes "P {}" 
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   774
  assumes "\<And>x S. finite S \<Longrightarrow> (\<And>y. y \<in> S \<Longrightarrow> f y \<le> f x) \<Longrightarrow> P S \<Longrightarrow> P (insert x S)"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   775
  shows "P S"
73102
87067698ae53 isabelle update_cartouches;
wenzelm
parents: 72384
diff changeset
   776
  using \<open>finite S\<close> 
70184
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   777
proof (induction rule: finite_psubset_induct)
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   778
  case (psubset A)
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   779
  {
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   780
    assume "A \<noteq> {}"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   781
    hence "f ` A \<noteq> {}" and "finite (f ` A)"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   782
      using psubset finite_image_iff by simp+ 
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   783
    then obtain a where "f a = Max (f ` A)" and "a \<in> A"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   784
      by (metis Max_in[of "f ` A"] imageE)
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   785
    then have "P (A - {a})"
82664
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 80934
diff changeset
   786
      using psubset(2) [of \<open>A - {a}\<close>] by auto
70184
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   787
    moreover 
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   788
    have "\<And>y. y \<in> A \<Longrightarrow> f y \<le> f a"
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   789
      using \<open>f a = Max (f ` A)\<close> \<open>finite (f ` A)\<close> by simp
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   790
    ultimately
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   791
    have ?case
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   792
      by (metis \<open>a \<in> A\<close> DiffD1 insert_Diff assms(3) finite_Diff psubset.hyps)
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   793
  }
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   794
  thus ?case
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   795
    using assms(2) by blast
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   796
qed
a7aba6db79a1 added lemma
nipkow
parents: 69593
diff changeset
   797
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   798
lemma Least_Min:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   799
  assumes "finite {a. P a}" and "\<exists>a. P a"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   800
  shows "(LEAST a. P a) = Min {a. P a}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   801
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   802
  { fix A :: "'a set"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   803
    assume A: "finite A" "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   804
    have "(LEAST a. a \<in> A) = Min A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   805
    using A proof (induct A rule: finite_ne_induct)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   806
      case singleton show ?case by (rule Least_equality) simp_all
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   807
    next
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   808
      case (insert a A)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   809
      have "(LEAST b. b = a \<or> b \<in> A) = min a (LEAST a. a \<in> A)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   810
        by (auto intro!: Least_equality simp add: min_def not_le Min_le_iff insert.hyps dest!: less_imp_le)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   811
      with insert show ?case by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   812
    qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   813
  } from this [of "{a. P a}"] assms show ?thesis by simp
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   814
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   815
82688
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   816
lemma Greatest_Max:
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   817
  assumes "finite {a. P a}" and "\<exists>a. P a"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   818
  shows "(GREATEST a. P a) = Max {a. P a}"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   819
proof -
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   820
  { fix A :: "'a set"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   821
    assume A: "finite A" "A \<noteq> {}"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   822
    have "(GREATEST a. a \<in> A) = Max A"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   823
    using A proof (induct A rule: finite_ne_induct)
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   824
      case singleton show ?case by (rule Greatest_equality) simp_all
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   825
    next
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   826
      case (insert a A)
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   827
      have "(GREATEST b. b = a \<or> b \<in> A) = max a (GREATEST a. a \<in> A)"
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   828
        by (auto intro!: Greatest_equality simp add: max_def not_le insert.hyps)
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   829
      with insert show ?case by simp
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   830
    qed
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   831
  } from this [of "{a. P a}"] assms show ?thesis by simp
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   832
qed
b391142bd2d2 prefer already existing operation to calculate minimum
haftmann
parents: 82664
diff changeset
   833
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   834
lemma infinite_growing:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   835
  assumes "X \<noteq> {}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   836
  assumes *: "\<And>x. x \<in> X \<Longrightarrow> \<exists>y\<in>X. y > x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   837
  shows "\<not> finite X"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   838
proof
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   839
  assume "finite X"
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 59000
diff changeset
   840
  with \<open>X \<noteq> {}\<close> have "Max X \<in> X" "\<forall>x\<in>X. x \<le> Max X"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   841
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   842
  with *[of "Max X"] show False
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   843
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   844
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58889
diff changeset
   845
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   846
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   847
73295
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   848
lemma sum_le_card_Max: "finite A \<Longrightarrow> sum f A \<le> card A * Max (f ` A)"
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   849
using sum_bounded_above[of A f "Max (f ` A)"] by simp
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   850
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   851
lemma card_Min_le_sum: "finite A \<Longrightarrow> card A * Min (f ` A) \<le> sum f A"
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   852
using sum_bounded_below[of A "Min (f ` A)" f] by simp
6c4c37a3ebec dropped obscure FIXME
haftmann
parents: 73221
diff changeset
   853
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   854
context linordered_ab_semigroup_add
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   855
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   856
68788
d4426a23832e tuned lemmas
nipkow
parents: 68779
diff changeset
   857
lemma Min_add_commute:
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   858
  fixes k
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   859
  assumes "finite S" and "S \<noteq> {}"
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   860
  shows "Min ((\<lambda>x. f x + k) ` S) = Min(f ` S) + k"
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   861
proof -
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   862
  have m: "\<And>x y. min x y + k = min (x+k) (y+k)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   863
    by (simp add: min_def order.antisym add_right_mono)
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   864
  have "(\<lambda>x. f x + k) ` S = (\<lambda>y. y + k) ` (f ` S)" by auto
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   865
  also have "Min \<dots> = Min (f ` S) + k"
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   866
    using assms hom_Min_commute [of "\<lambda>y. y+k" "f ` S", OF m, symmetric] by simp
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   867
  finally show ?thesis by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   868
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   869
68779
78d9b1783378 copied but not adapted
nipkow
parents: 67969
diff changeset
   870
lemma Max_add_commute:
78d9b1783378 copied but not adapted
nipkow
parents: 67969
diff changeset
   871
  fixes k
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   872
  assumes "finite S" and "S \<noteq> {}"
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   873
  shows "Max ((\<lambda>x. f x + k) ` S) = Max(f ` S) + k"
68779
78d9b1783378 copied but not adapted
nipkow
parents: 67969
diff changeset
   874
proof -
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   875
  have m: "\<And>x y. max x y + k = max (x+k) (y+k)"
73411
1f1366966296 avoid name clash
haftmann
parents: 73326
diff changeset
   876
    by (simp add: max_def order.antisym add_right_mono)
68793
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   877
  have "(\<lambda>x. f x + k) ` S = (\<lambda>y. y + k) ` (f ` S)" by auto
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   878
  also have "Max \<dots> = Max (f ` S) + k"
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   879
    using assms hom_Max_commute [of "\<lambda>y. y+k" "f ` S", OF m, symmetric] by simp
462226db648a tuned lemmas
nipkow
parents: 68788
diff changeset
   880
  finally show ?thesis by simp
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   881
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   882
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   883
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   884
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   885
context linordered_ab_group_add
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   886
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   887
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   888
lemma minus_Max_eq_Min [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   889
  "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Max S = Min (uminus ` S)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   890
  by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   891
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   892
lemma minus_Min_eq_Max [simp]:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   893
  "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - Min S = Max (uminus ` S)"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   894
  by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   895
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   896
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   897
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   898
context complete_linorder
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   899
begin
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   900
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   901
lemma Min_Inf:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   902
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   903
  shows "Min A = Inf A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   904
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   905
  from assms obtain b B where "A = insert b B" and "finite B" by auto
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   906
  then show ?thesis
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   907
    by (simp add: Min.eq_fold complete_linorder_inf_min [symmetric] inf_Inf_fold_inf inf.commute [of b])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   908
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   909
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   910
lemma Max_Sup:
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   911
  assumes "finite A" and "A \<noteq> {}"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   912
  shows "Max A = Sup A"
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   913
proof -
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   914
  from assms obtain b B where "A = insert b B" and "finite B" by auto
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   915
  then show ?thesis
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   916
    by (simp add: Max.eq_fold complete_linorder_sup_max [symmetric] sup_Sup_fold_sup sup.commute [of b])
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   917
qed
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   918
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   919
end
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
   920
73326
7a88313895d5 dissolve theory with duplicated name from afp
haftmann
parents: 73295
diff changeset
   921
lemma disjnt_ge_max: \<^marker>\<open>contributor \<open>Lars Hupel\<close>\<close>
7a88313895d5 dissolve theory with duplicated name from afp
haftmann
parents: 73295
diff changeset
   922
  \<open>disjnt X Y\<close> if \<open>finite Y\<close> \<open>\<And>x. x \<in> X \<Longrightarrow> x > Max Y\<close>
7a88313895d5 dissolve theory with duplicated name from afp
haftmann
parents: 73295
diff changeset
   923
  using that by (auto simp add: disjnt_def) (use Max_less_iff in blast)
7a88313895d5 dissolve theory with duplicated name from afp
haftmann
parents: 73295
diff changeset
   924
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
   925
82824
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   926
subsection \<open>An aside: code generation for \<open>LEAST\<close> and \<open>GREATEST\<close>\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   927
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   928
context
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   929
begin
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   930
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   931
qualified definition Least :: \<open>'a::linorder set \<Rightarrow> 'a\<close> \<comment> \<open>only for code generation\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   932
  where Least_eq [code_abbrev, simp]: \<open>Least S = (LEAST x. x \<in> S)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   933
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   934
qualified lemma Least_filter_eq [code_abbrev]:
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   935
  \<open>Least (Set.filter P S) = (LEAST x. x \<in> S \<and> P x)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   936
  by simp
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   937
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   938
qualified definition Least_abort :: \<open>'a set \<Rightarrow> 'a::linorder\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   939
  where \<open>Least_abort = Least\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   940
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82824
diff changeset
   941
qualified lemma Least_code [code abort: Lattices_Big.Least_abort, code]:
82824
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   942
  \<open>Least A = (if finite A \<longrightarrow> Set.is_empty A then Least_abort A else Min A)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   943
  using Least_Min [of \<open>\<lambda>x. x \<in> A\<close>] by (auto simp add: Least_abort_def)
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   944
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   945
qualified definition Greatest :: \<open>'a::linorder set \<Rightarrow> 'a\<close> \<comment> \<open>only for code generation\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   946
  where Greatest_eq [code_abbrev, simp]: \<open>Greatest S = (GREATEST x. x \<in> S)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   947
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   948
qualified lemma Greatest_filter_eq [code_abbrev]:
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   949
  \<open>Greatest (Set.filter P S) = (GREATEST x. x \<in> S \<and> P x)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   950
  by simp
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   951
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   952
qualified definition Greatest_abort :: \<open>'a set \<Rightarrow> 'a::linorder\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   953
  where \<open>Greatest_abort = Greatest\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   954
82858
52cf8f3f1652 always use proper context when parsing constants
haftmann
parents: 82824
diff changeset
   955
qualified lemma Greatest_code [code abort: Lattices_Big.Greatest_abort, code]:
82824
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   956
  \<open>Greatest A = (if finite A \<longrightarrow> Set.is_empty A then Greatest_abort A else Max A)\<close>
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   957
  using Greatest_Max [of \<open>\<lambda>x. x \<in> A\<close>] by (auto simp add: Greatest_abort_def)
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   958
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   959
end
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   960
7ddae44464d4 moved to more appropriate theory
haftmann
parents: 82688
diff changeset
   961
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
   962
subsection \<open>Arg Min\<close>
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
   963
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   964
context ord
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   965
begin
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   966
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   967
definition is_arg_min :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" where
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   968
"is_arg_min f P x = (P x \<and> \<not>(\<exists>y. P y \<and> f y < f x))"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   969
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   970
definition arg_min :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'b" where
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   971
"arg_min f P = (SOME x. is_arg_min f P x)"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   972
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   973
definition arg_min_on :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'b" where
67169
1fabca1c2199 made arg_min_on definition
nipkow
parents: 67036
diff changeset
   974
"arg_min_on f S = arg_min f (\<lambda>x. x \<in> S)"
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   975
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   976
end
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   977
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   978
syntax
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
   979
  "_arg_min" :: "('b \<Rightarrow> 'a) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'b"
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
   980
    (\<open>(\<open>indent=3 notation=\<open>binder ARG_MIN\<close>\<close>ARG'_MIN _ _./ _)\<close> [1000, 0, 10] 10)
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   981
syntax_consts
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
   982
  "_arg_min" \<rightleftharpoons> arg_min
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   983
translations
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   984
  "ARG_MIN f x. P" \<rightleftharpoons> "CONST arg_min f (\<lambda>x. P)"
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   985
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   986
lemma is_arg_min_linorder: fixes f :: "'a \<Rightarrow> 'b :: linorder"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   987
shows "is_arg_min f P x = (P x \<and> (\<forall>y. P y \<longrightarrow> f x \<le> f y))"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
   988
by(auto simp add: is_arg_min_def)
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
   989
67566
c555f1778dd8 added lemma
nipkow
parents: 67525
diff changeset
   990
lemma is_arg_min_antimono: fixes f :: "'a \<Rightarrow> ('b::order)"
c555f1778dd8 added lemma
nipkow
parents: 67525
diff changeset
   991
shows "\<lbrakk> is_arg_min f P x; f y \<le> f x; P y \<rbrakk> \<Longrightarrow> is_arg_min f P y"
c555f1778dd8 added lemma
nipkow
parents: 67525
diff changeset
   992
by (simp add: order.order_iff_strict is_arg_min_def)
c555f1778dd8 added lemma
nipkow
parents: 67525
diff changeset
   993
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   994
lemma arg_minI:
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   995
  "\<lbrakk> P x;
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   996
    \<And>y. P y \<Longrightarrow> \<not> f y < f x;
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   997
    \<And>x. \<lbrakk> P x; \<forall>y. P y \<longrightarrow> \<not> f y < f x \<rbrakk> \<Longrightarrow> Q x \<rbrakk>
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
   998
  \<Longrightarrow> Q (arg_min f P)"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
   999
  unfolding arg_min_def is_arg_min_def
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1000
  by (blast intro!: someI2_ex)
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1001
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1002
lemma arg_min_equality:
65952
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1003
  "\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f k \<le> f x \<rbrakk> \<Longrightarrow> f (arg_min f P) = f k"
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1004
  for f :: "_ \<Rightarrow> 'a::order"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1005
  by (rule arg_minI; force simp: not_less less_le_not_le)
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1006
65952
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1007
lemma wf_linord_ex_has_least:
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1008
  "\<lbrakk> wf r; \<forall>x y. (x, y) \<in> r\<^sup>+ \<longleftrightarrow> (y, x) \<notin> r\<^sup>*; P k \<rbrakk>
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1009
   \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> (m x, m y) \<in> r\<^sup>*)"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1010
  by (force dest!:  wf_trancl [THEN wf_eq_minimal [THEN iffD1, THEN spec], where x = "m ` Collect P"])
65952
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1011
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1012
lemma ex_has_least_nat: "P k \<Longrightarrow> \<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> m x \<le> m y)"
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1013
  for m :: "'a \<Rightarrow> nat"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1014
  unfolding pred_nat_trancl_eq_le [symmetric]
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1015
  apply (rule wf_pred_nat [THEN wf_linord_ex_has_least])
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1016
   apply (simp add: less_eq linorder_not_le pred_nat_trancl_eq_le)
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1017
  by assumption
65952
dec96cb3fbe0 removed LeastM; is now arg_min
nipkow
parents: 65951
diff changeset
  1018
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1019
lemma arg_min_nat_lemma:
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1020
  "P k \<Longrightarrow> P(arg_min m P) \<and> (\<forall>y. P y \<longrightarrow> m (arg_min m P) \<le> m y)"
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1021
  for m :: "'a \<Rightarrow> nat"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1022
  unfolding arg_min_def is_arg_min_linorder
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1023
  apply (rule someI_ex)
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1024
  apply (erule ex_has_least_nat)
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1025
  done
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1026
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1027
lemmas arg_min_natI = arg_min_nat_lemma [THEN conjunct1]
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1028
65951
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1029
lemma is_arg_min_arg_min_nat: fixes m :: "'a \<Rightarrow> nat"
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1030
shows "P x \<Longrightarrow> is_arg_min m P (arg_min m P)"
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1031
by (metis arg_min_nat_lemma is_arg_min_linorder)
32b3feb6f965 more arg_min
nipkow
parents: 65842
diff changeset
  1032
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1033
lemma arg_min_nat_le: "P x \<Longrightarrow> m (arg_min m P) \<le> m x"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1034
  for m :: "'a \<Rightarrow> nat"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1035
  by (rule arg_min_nat_lemma [THEN conjunct2, THEN spec, THEN mp])
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1036
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1037
lemma ex_min_if_finite:
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1038
  "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>m\<in>S. \<not>(\<exists>x\<in>S. x < (m::'a::order))"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1039
  by(induction rule: finite.induct) (auto intro: order.strict_trans)
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1040
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1041
lemma ex_is_arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1042
  shows "\<lbrakk> finite S; S \<noteq> {} \<rbrakk> \<Longrightarrow> \<exists>x. is_arg_min f (\<lambda>x. x \<in> S) x"
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1043
  unfolding is_arg_min_def
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1044
  using ex_min_if_finite[of "f ` S"]
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1045
  by auto
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1046
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1047
lemma arg_min_SOME_Min:
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1048
  "finite S \<Longrightarrow> arg_min_on f S = (SOME y. y \<in> S \<and> f y = Min(f ` S))"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1049
  unfolding arg_min_on_def arg_min_def is_arg_min_linorder
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1050
  apply(rule arg_cong[where f = Eps])
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1051
  apply (auto simp: fun_eq_iff intro: Min_eqI[symmetric])
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1052
  done
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1053
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1054
lemma arg_min_if_finite: fixes f :: "'a \<Rightarrow> 'b :: order"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1055
  assumes "finite S" "S \<noteq> {}"
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1056
  shows  "arg_min_on f S \<in> S" and "\<not>(\<exists>x\<in>S. f x < f (arg_min_on f S))"
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1057
  using ex_is_arg_min_if_finite[OF assms, of f]
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1058
  unfolding arg_min_on_def arg_min_def is_arg_min_def
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1059
  by(auto dest!: someI_ex)
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1060
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1061
lemma arg_min_least: fixes f :: "'a \<Rightarrow> 'b :: linorder"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1062
  shows "\<lbrakk> finite S;  S \<noteq> {};  y \<in> S \<rbrakk> \<Longrightarrow> f(arg_min_on f S) \<le> f y"
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1063
  by(simp add: arg_min_SOME_Min inv_into_def2[symmetric] f_inv_into_f)
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1064
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1065
lemma arg_min_inj_eq: fixes f :: "'a \<Rightarrow> 'b :: order"
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1066
shows "\<lbrakk> inj_on f {x. P x}; P a; \<forall>y. P y \<longrightarrow> f a \<le> f y \<rbrakk> \<Longrightarrow> arg_min f P = a"
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1067
apply(simp add: arg_min_def is_arg_min_def)
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1068
apply(rule someI2[of _ a])
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1069
 apply (simp add: less_le_not_le)
65842
42420ae446a2 more arg_min
nipkow
parents: 65817
diff changeset
  1070
by (metis inj_on_eq_iff less_le mem_Collect_eq)
65817
8ee1799fb076 added function arg_min
nipkow
parents: 63915
diff changeset
  1071
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1072
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1073
subsection \<open>Arg Max\<close>
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1074
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1075
context ord
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1076
begin
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1077
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1078
definition is_arg_max :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'b \<Rightarrow> bool" where
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1079
"is_arg_max f P x = (P x \<and> \<not>(\<exists>y. P y \<and> f y > f x))"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1080
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1081
definition arg_max :: "('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'b" where
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1082
"arg_max f P = (SOME x. is_arg_max f P x)"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1083
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1084
definition arg_max_on :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'b" where
67169
1fabca1c2199 made arg_min_on definition
nipkow
parents: 67036
diff changeset
  1085
"arg_max_on f S = arg_max f (\<lambda>x. x \<in> S)"
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1086
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1087
end
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1088
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1089
syntax
72384
b037517c815b added lemmas; internalized defn in class
nipkow
parents: 70184
diff changeset
  1090
  "_arg_max" :: "('b \<Rightarrow> 'a) \<Rightarrow> pttrn \<Rightarrow> bool \<Rightarrow> 'a"
80934
8e72f55295fd more inner syntax markup: HOL;
wenzelm
parents: 80932
diff changeset
  1091
    (\<open>(\<open>indent=3 notation=\<open>binder ARG_MAX\<close>\<close>ARG'_MAX _ _./ _)\<close> [1000, 0, 10] 10)
80760
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
  1092
syntax_consts
be8c0e039a5e more markup for syntax consts;
wenzelm
parents: 80175
diff changeset
  1093
  "_arg_max" \<rightleftharpoons> arg_max
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1094
translations
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1095
  "ARG_MAX f x. P" \<rightleftharpoons> "CONST arg_max f (\<lambda>x. P)"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1096
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1097
lemma is_arg_max_linorder: fixes f :: "'a \<Rightarrow> 'b :: linorder"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1098
shows "is_arg_max f P x = (P x \<and> (\<forall>y. P y \<longrightarrow> f x \<ge> f y))"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1099
by(auto simp add: is_arg_max_def)
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1100
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1101
lemma arg_maxI:
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1102
  "P x \<Longrightarrow>
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1103
    (\<And>y. P y \<Longrightarrow> \<not> f y > f x) \<Longrightarrow>
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1104
    (\<And>x. P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> \<not> f y > f x \<Longrightarrow> Q x) \<Longrightarrow>
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1105
    Q (arg_max f P)"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1106
  unfolding arg_max_def is_arg_max_def
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1107
  by (blast intro!: someI2_ex elim: )
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1108
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1109
lemma arg_max_equality:
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1110
  "\<lbrakk> P k; \<And>x. P x \<Longrightarrow> f x \<le> f k \<rbrakk> \<Longrightarrow> f (arg_max f P) = f k"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1111
  for f :: "_ \<Rightarrow> 'a::order"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1112
apply (rule arg_maxI [where f = f])
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1113
  apply assumption
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1114
 apply (simp add: less_le_not_le)
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1115
by (metis le_less)
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1116
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1117
lemma ex_has_greatest_nat_lemma:
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1118
  "P k \<Longrightarrow> \<forall>x. P x \<longrightarrow> (\<exists>y. P y \<and> \<not> f y \<le> f x) \<Longrightarrow> \<exists>y. P y \<and> \<not> f y < f k + n"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1119
  for f :: "'a \<Rightarrow> nat"
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1120
  by (induct n) (force simp: le_Suc_eq)+
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1121
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1122
lemma ex_has_greatest_nat:
75669
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1123
  assumes "P k"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1124
    and "\<forall>y. P y \<longrightarrow> (f:: 'a \<Rightarrow> nat) y < b"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1125
shows "\<exists>x. P x \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f x)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1126
proof (rule ccontr)
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1127
  assume "\<nexists>x. P x \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f x)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1128
  then have "\<forall>x. P x \<longrightarrow> (\<exists>y. P y \<and> \<not> f y \<le> f x)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1129
    by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1130
  then have "\<exists>y. P y \<and> \<not> f y < f k + (b - f k)"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1131
    using assms ex_has_greatest_nat_lemma[of P k f "b - f k"]
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1132
    by blast
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1133
  then show "False"
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1134
    using assms by auto
43f5dfb7fa35 tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents: 74979
diff changeset
  1135
qed
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1136
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1137
lemma arg_max_nat_lemma:
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1138
  "\<lbrakk> P k;  \<forall>y. P y \<longrightarrow> f y < b \<rbrakk>
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1139
  \<Longrightarrow> P (arg_max f P) \<and> (\<forall>y. P y \<longrightarrow> f y \<le> f (arg_max f P))"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1140
  for f :: "'a \<Rightarrow> nat"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1141
  unfolding arg_max_def is_arg_max_linorder
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1142
  by (rule someI_ex) (metis ex_has_greatest_nat)
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1143
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1144
lemmas arg_max_natI = arg_max_nat_lemma [THEN conjunct1]
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1145
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1146
lemma arg_max_nat_le: "P x \<Longrightarrow> \<forall>y. P y \<longrightarrow> f y < b \<Longrightarrow> f x \<le> f (arg_max f P)"
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1147
  for f :: "'a \<Rightarrow> nat"
80769
77f7aa898ced A bit more tidying
paulson <lp15@cam.ac.uk>
parents: 80175
diff changeset
  1148
  using arg_max_nat_lemma by metis
65954
431024edc9cf introduced arg_max
nipkow
parents: 65952
diff changeset
  1149
54744
1e7f2d296e19 more algebraic terminology for theories about big operators
haftmann
parents:
diff changeset
  1150
end