Theory Stone_Algebras.Lattice_Basics
section ‹Lattice Basics›
text ‹
This theory provides notations, basic definitions and facts of lattice-related structures used throughout the subsequent development.
›
theory Lattice_Basics
imports Main
begin
subsection ‹General Facts and Notations›
text ‹
The following results extend basic Isabelle/HOL facts.
›
lemma imp_as_conj:
  assumes "P x ⟹ Q x"
  shows "P x ∧ Q x ⟷ P x"
  using assms by auto
lemma if_distrib_2:
  "f (if c then x else y) (if c then z else w) = (if c then f x z else f y w)"
  by simp
lemma left_invertible_inj:
  "(∀x . g (f x) = x) ⟹ inj f"
  by (metis injI)
lemma invertible_bij:
  assumes "∀x . g (f x) = x"
      and "∀y . f (g y) = y"
    shows "bij f"
  by (metis assms bijI')
lemma finite_ne_subset_induct [consumes 3, case_names singleton insert]:
  assumes "finite F"
      and "F ≠ {}"
      and "F ⊆ S"
      and singleton: "⋀x . P {x}"
      and insert: "⋀x F . finite F ⟹ F ≠ {} ⟹ F ⊆ S ⟹ x ∈ S ⟹ x ∉ F ⟹ P F ⟹ P (insert x F)"
    shows "P F"
  using assms(1-3)
  apply (induct rule: finite_ne_induct)
  apply (simp add: singleton)
  by (simp add: insert)
lemma finite_set_of_finite_funs_pred:
  assumes "finite { x::'a . True }"
      and "finite { y::'b . P y }"
    shows "finite { f . (∀x::'a . P (f x)) }"
  using assms finite_set_of_finite_funs by force
text ‹
We use the following notations for the join, meet and complement operations.
Changing the precedence of the unary complement allows us to write terms like ‹--x› instead of ‹-(-x)›.
›
context sup
begin
notation sup (infixl ‹⊔› 65)
definition additive :: "('a ⇒ 'a) ⇒ bool"
  where "additive f ≡ ∀x y . f (x ⊔ y) = f x ⊔ f y"
end
context inf
begin
notation inf (infixl ‹⊓› 67)
end
context uminus
begin
unbundle no uminus_syntax
notation uminus (‹(‹open_block notation=‹prefix -››- _)› [80] 80)
end
subsection ‹Orders›
text ‹
We use the following definition of monotonicity for operations defined in classes.
The standard ‹mono› places a sort constraint on the target type.
We also give basic properties of Galois connections and lift orders to functions.
›
context ord
begin
definition isotone :: "('a ⇒ 'a) ⇒ bool"
  where "isotone f ≡ ∀x y . x ≤ y ⟶ f x ≤ f y"
definition galois :: "('a ⇒ 'a) ⇒ ('a ⇒ 'a) ⇒ bool"
  where "galois l u ≡ ∀x y . l x ≤ y ⟷ x ≤ u y"
definition lifted_less_eq :: "('a ⇒ 'a) ⇒ ('a ⇒ 'a) ⇒ bool" (‹(_ ≤≤ _)› [51, 51] 50)
  where "f ≤≤ g ≡ ∀x . f x ≤ g x"
end
context order
begin
lemma order_lesseq_imp:
  "(∀z . x ≤ z ⟶ y ≤ z) ⟷ y ≤ x"
  using order_trans by blast
lemma galois_char:
  "galois l u ⟷ (∀x . x ≤ u (l x)) ∧ (∀x . l (u x) ≤ x) ∧ isotone l ∧ isotone u"
  apply (rule iffI)
  apply (metis (full_types) galois_def isotone_def order_refl order_trans)
  using galois_def isotone_def order_trans by blast
lemma galois_closure:
  "galois l u ⟹ l x = l (u (l x)) ∧ u x = u (l (u x))"
  by (simp add: galois_char isotone_def order.antisym)
lemma lifted_reflexive:
  "f = g ⟹ f ≤≤ g"
  by (simp add: lifted_less_eq_def)
lemma lifted_transitive:
  "f ≤≤ g ⟹ g ≤≤ h ⟹ f ≤≤ h"
  using lifted_less_eq_def order_trans by blast
lemma lifted_antisymmetric:
  "f ≤≤ g ⟹ g ≤≤ f ⟹ f = g"
  by (rule ext, rule order.antisym) (simp_all add: lifted_less_eq_def)
text ‹
If the image of a finite non-empty set under ‹f› is a totally ordered, there is an element that minimises the value of ‹f›.
›
lemma finite_set_minimal:
  assumes "finite s"
      and "s ≠ {}"
      and "∀x∈s . ∀y∈s . f x ≤ f y ∨ f y ≤ f x"
    shows "∃m∈s . ∀z∈s . f m ≤ f z"
  apply (rule finite_ne_subset_induct[where S=s])
  apply (rule assms(1))
  apply (rule assms(2))
  apply simp
  apply simp
  by (metis assms(3) insert_iff order_trans subsetD)
end
subsection ‹Semilattices›
text ‹
The following are basic facts in semilattices.
›
context semilattice_sup
begin
lemma sup_left_isotone:
  "x ≤ y ⟹ x ⊔ z ≤ y ⊔ z"
  using sup.mono by blast
lemma sup_right_isotone:
  "x ≤ y ⟹ z ⊔ x ≤ z ⊔ y"
  using sup.mono by blast
lemma sup_left_divisibility:
  "x ≤ y ⟷ (∃z . x ⊔ z = y)"
  using sup.absorb2 sup.cobounded1 by blast
lemma sup_right_divisibility:
  "x ≤ y ⟷ (∃z . z ⊔ x = y)"
  by (metis sup.cobounded2 sup.orderE)
lemma sup_same_context:
  "x ≤ y ⊔ z ⟹ y ≤ x ⊔ z ⟹ x ⊔ z = y ⊔ z"
  by (simp add: le_iff_sup sup_left_commute)
lemma sup_relative_same_increasing:
  "x ≤ y ⟹ x ⊔ z = x ⊔ w ⟹ y ⊔ z = y ⊔ w"
  using sup.assoc sup_right_divisibility by auto
end
text ‹
Every bounded semilattice is a commutative monoid.
Finite sums defined in commutative monoids are available via the following sublocale.
›
context bounded_semilattice_sup_bot
begin