| author | haftmann | 
| Sat, 24 Dec 2011 16:14:58 +0100 | |
| changeset 45980 | af59825c40cf | 
| parent 45903 | 02dd9319dcb7 | 
| child 46066 | e81411bfa7ef | 
| permissions | -rw-r--r-- | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 1 | theory Complete_Lattice_ix | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 2 | imports Main | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 3 | begin | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 4 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 5 | text{* A complete lattice is an ordered type where every set of elements has
 | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 6 | a greatest lower (and thus also a leats upper) bound. Sets are the | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 7 | prototypical complete lattice where the greatest lower bound is | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 8 | intersection. Sometimes that set of all elements of a type is not a complete | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 9 | lattice although all elements of the same shape form a complete lattice, for | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 10 | example lists of the same length, where the list elements come from a | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 11 | complete lattice. We will have exactly this situation with annotated | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 12 | commands. This theory introduces a slightly generalised version of complete | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 13 | lattices where elements have an ``index'' and only the set of elements with | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 14 | the same index form a complete lattice; the type as a whole is a disjoint | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 15 | union of complete lattices. Because sets are not types, this requires a | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 16 | special treatment. *} | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 17 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 18 | locale Complete_Lattice_ix = | 
| 45903 | 19 | fixes L :: "'i \<Rightarrow> 'a::order set" | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 20 | and Glb :: "'i \<Rightarrow> 'a set \<Rightarrow> 'a" | 
| 45903 | 21 | assumes Glb_lower: "A \<subseteq> L i \<Longrightarrow> a \<in> A \<Longrightarrow> (Glb i A) \<le> a" | 
| 22 | and Glb_greatest: "b : L i \<Longrightarrow> \<forall>a\<in>A. b \<le> a \<Longrightarrow> b \<le> (Glb i A)" | |
| 23 | and Glb_in_L: "A \<subseteq> L i \<Longrightarrow> Glb i A : L i" | |
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 24 | begin | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 25 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 26 | definition lfp :: "'i \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 45903 | 27 | "lfp i f = Glb i {a : L i. f a \<le> a}"
 | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 28 | |
| 45903 | 29 | lemma index_lfp: "lfp i f : L i" | 
| 30 | by(auto simp: lfp_def intro: Glb_in_L) | |
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 31 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 32 | lemma lfp_lowerbound: | 
| 45903 | 33 | "\<lbrakk> a : L i; f a \<le> a \<rbrakk> \<Longrightarrow> lfp i f \<le> a" | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 34 | by (auto simp add: lfp_def intro: Glb_lower) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 35 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 36 | lemma lfp_greatest: | 
| 45903 | 37 | "\<lbrakk> a : L i; \<And>u. \<lbrakk> u : L i; f u \<le> u\<rbrakk> \<Longrightarrow> a \<le> u \<rbrakk> \<Longrightarrow> a \<le> lfp i f" | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 38 | by (auto simp add: lfp_def intro: Glb_greatest) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 39 | |
| 45903 | 40 | lemma lfp_unfold: assumes "\<And>x i. f x : L i \<longleftrightarrow> x : L i" | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 41 | and mono: "mono f" shows "lfp i f = f (lfp i f)" | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 42 | proof- | 
| 45903 | 43 | note assms(1)[simp] index_lfp[simp] | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 44 | have 1: "f (lfp i f) \<le> lfp i f" | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 45 | apply(rule lfp_greatest) | 
| 45903 | 46 | apply simp | 
| 45623 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 47 | by (blast intro: lfp_lowerbound monoD[OF mono] order_trans) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 48 | have "lfp i f \<le> f (lfp i f)" | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 49 | by (fastforce intro: 1 monoD[OF mono] lfp_lowerbound) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 50 | with 1 show ?thesis by(blast intro: order_antisym) | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 51 | qed | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 52 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 53 | end | 
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 54 | |
| 
f682f3f7b726
Abstract interpretation is now based uniformly on annotated programs,
 nipkow parents: diff
changeset | 55 | end |