src/HOL/IMP/Complete_Lattice.thy
author haftmann
Fri, 24 Aug 2018 20:22:14 +0000
changeset 68802 3974935e0252
parent 68778 4566bac4517d
permissions -rw-r--r--
some modernization of notation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68778
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     1
(* Author: Tobias Nipkow *)
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     2
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 49487
diff changeset
     3
section "Abstract Interpretation"
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 48759
diff changeset
     4
68778
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     5
subsection "Complete Lattice"
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     6
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
     7
theory Complete_Lattice
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
     8
imports Main
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
     9
begin
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    10
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    11
locale Complete_Lattice =
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    12
fixes L :: "'a::order set" and Glb :: "'a set \<Rightarrow> 'a"
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    13
assumes Glb_lower: "A \<subseteq> L \<Longrightarrow> a \<in> A \<Longrightarrow> Glb A \<le> a"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    14
and Glb_greatest: "b \<in> L \<Longrightarrow> \<forall>a\<in>A. b \<le> a \<Longrightarrow> b \<le> Glb A"
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    15
and Glb_in_L: "A \<subseteq> L \<Longrightarrow> Glb A \<in> L"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    16
begin
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    17
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    18
definition lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    19
"lfp f = Glb {a : L. f a \<le> a}"
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    20
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    21
lemma index_lfp: "lfp f \<in> L"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    22
by(auto simp: lfp_def intro: Glb_in_L)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    23
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    24
lemma lfp_lowerbound:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    25
  "\<lbrakk> a \<in> L;  f a \<le> a \<rbrakk> \<Longrightarrow> lfp f \<le> a"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    26
by (auto simp add: lfp_def intro: Glb_lower)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    27
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    28
lemma lfp_greatest:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    29
  "\<lbrakk> a \<in> L;  \<And>u. \<lbrakk> u \<in> L; f u \<le> u\<rbrakk> \<Longrightarrow> a \<le> u \<rbrakk> \<Longrightarrow> a \<le> lfp f"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    30
by (auto simp add: lfp_def intro: Glb_greatest)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    31
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 58889
diff changeset
    32
lemma lfp_unfold: assumes "\<And>x. f x \<in> L \<longleftrightarrow> x \<in> L"
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    33
and mono: "mono f" shows "lfp f = f (lfp f)"
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    34
proof-
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    35
  note assms(1)[simp] index_lfp[simp]
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    36
  have 1: "f (lfp f) \<le> lfp f"
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    37
    apply(rule lfp_greatest)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    38
    apply simp
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    39
    by (blast intro: lfp_lowerbound monoD[OF mono] order_trans)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    40
  have "lfp f \<le> f (lfp f)"
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    41
    by (fastforce intro: 1 monoD[OF mono] lfp_lowerbound)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    42
  with 1 show ?thesis by(blast intro: order_antisym)
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    43
qed
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    44
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    45
end
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    46
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    47
end
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents:
diff changeset
    48