Theory Sup_Lattice

(* 
  Title: Sup-Lattices and Other Simplifications
  Author: Georg Struth 
  Maintainer: Georg Struth <g.struth@sheffield.ac.uk> 
*)

section ‹Sup-Lattices and Other Simplifications›

theory Sup_Lattice
  imports Main
begin

unbundle lattice_syntax

text ‹Some definitions for orderings and lattices in Isabelle could be simpler. The strict order in 
in ord could be defined instead of being axiomatised. The function mono could have been defined on ord
and not on order---even on a general (di)graph it serves as a morphism. In complete lattices, the 
supremum---and dually the infimum---suffices to define the other operations (in the Isabelle/HOL-definition
infimum, binary supremum and infimum, bottom and top element are axiomatised). This not only increases
the number of proof obligations in subclass or sublocale statements, instantiations or interpretations,
it also complicates situations where suprema are presented faithfully, e.g. mapped onto suprema in 
some subalgebra, whereas infima in the subalgebra are different from those in the super-structure.›


text ‹It would be even nicer to use a class less-eq which dispenses with the strict order symbol
in ord. Then one would not have to redefine this symbol in all instantiations or interpretations.
At least, it does not carry any proof obligations.›

context ord
begin

text ‹ub-set yields the set of all upper bounds of a set; lb-set the set of all lower bounds.›

definition ub_set :: "'a set  'a set" where
  "ub_set X = {y. x  X. x  y}"

definition lb_set :: "'a set  'a set" where
  "lb_set X = {y. x  X. y  x}"

end

definition ord_pres :: "('a::ord  'b::ord)  bool" where
 "ord_pres f = (x y. x  y  f x  f y)"

lemma ord_pres_mono: 
  fixes f :: "'a::order  'b::order"
  shows "mono f = ord_pres f"
  by (simp add: mono_def ord_pres_def)

class preorder_lean = ord +
  assumes preorder_refl: "x  x"
  and preorder_trans: "x  y  y  z  x  z"

begin

definition le :: "'a  'a  bool" where
  "le x y = (x  y  ¬ (x  y))"

end

sublocale preorder_lean  prel: preorder "(≤)" le
  by (unfold_locales, auto simp add: le_def preorder_refl preorder_trans)
 
class order_lean = preorder_lean +
  assumes order_antisym: "x  y  x  y  x = y"

sublocale order_lean  posl: order "(≤)" le
  by (unfold_locales, simp add: order_antisym)

class Sup_lattice = order_lean + Sup +
  assumes Sups_upper: "x  X  x  X"
  and Sups_least: "(x. x  X  x  z)  X  z"

begin

definition Infs :: "'a set  'a" where
  "Infs X =  {y. x  X. y  x}"

definition sups :: "'a  'a  'a" where
  "sups x y = {x,y}"

definition infs :: "'a  'a  'a" where
  "infs x y = Infs{x,y}"

definition bots :: 'a where 
  "bots = {}"

definition tops :: 'a where
  "tops = Infs{}"

lemma Infs_prop: "Infs = Sup  lb_set"
  unfolding fun_eq_iff by (simp add: Infs_def prel.lb_set_def)

end

class Inf_lattice = order_lean + Inf +
  assumes Infi_lower: "x  X  X  x"
  and Infi_greatest: "(x. x  X  z  x)  z  X"

begin

definition Supi :: "'a set  'a" where
  "Supi X = {y. x  X. x  y}"

definition supi :: "'a  'a  'a" where
  "supi x y = Supi{x,y}"

definition infi :: "'a  'a  'a" where
  "infi x y = {x,y}"

definition boti :: 'a where 
  "boti = Supi{}"

definition topi :: 'a where
  "topi = {}"

lemma Supi_prop: "Supi = Inf  ub_set"
  unfolding fun_eq_iff by (simp add: Supi_def prel.ub_set_def)

end

sublocale Inf_lattice  ldual: Sup_lattice Inf "(≥)"
  rewrites "ldual.Infs = Supi"
  and "ldual.infs = supi"
  and "ldual.sups = infi"
  and "ldual.tops = boti"
  and "ldual.bots = topi"
proof-
  show "class.Sup_lattice Inf (≥)"
    by (unfold_locales, simp_all add: Infi_lower Infi_greatest preorder_trans)
  then interpret ldual: Sup_lattice Inf "(≥)".
  show a: "ldual.Infs = Supi"
    unfolding fun_eq_iff by (simp add: ldual.Infs_def Supi_def)
  show "ldual.infs = supi"
    unfolding fun_eq_iff by (simp add: a ldual.infs_def supi_def)
  show "ldual.sups = infi"
    unfolding fun_eq_iff by (simp add: ldual.sups_def infi_def) 
  show "ldual.tops = boti"
    by (simp add: a ldual.tops_def boti_def)
  show "ldual.bots = topi"
    by (simp add: ldual.bots_def topi_def)
qed

sublocale Sup_lattice  supclat: complete_lattice Infs Sup_class.Sup infs "(≤)" le sups bots tops
  apply unfold_locales
  unfolding Infs_def infs_def sups_def bots_def tops_def
  by (simp_all, auto intro: Sups_least, simp_all add: Sups_upper)
       
 sublocale Inf_lattice  infclat: complete_lattice Inf_class.Inf Supi infi "(≤)" le supi boti topi
  by (unfold_locales, simp_all add: ldual.Sups_upper ldual.Sups_least ldual.supclat.Inf_lower ldual.supclat.Inf_greatest)

end