src/HOL/IMP/Astate.thy
author nipkow
Fri, 02 Sep 2011 19:25:18 +0200
changeset 44656 22bbd0d1b943
child 44923 b80108b346a9
permissions -rw-r--r--
Added Abstract Interpretation theories
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     2
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     3
theory Astate
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     4
imports AbsInt0_fun
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     5
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     6
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     7
subsection "Abstract State with Computable Ordering"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     8
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     9
text{* A concrete type of state with computable @{text"\<sqsubseteq>"}: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    10
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    11
datatype 'a astate = FunDom "string \<Rightarrow> 'a" "string list"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    12
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    13
fun "fun" where "fun (FunDom f _) = f"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    14
fun dom where "dom (FunDom _ A) = A"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    15
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    16
definition "list S = [(x,fun S x). x \<leftarrow> dom S]"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    17
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    18
definition "lookup F x = (if x : set(dom F) then fun F x else Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    19
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    20
definition "update F x y =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    21
  FunDom ((fun F)(x:=y)) (if x : set(dom F) then dom F else x # dom F)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    22
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    23
lemma lookup_update: "lookup (update S x y) = (lookup S)(x:=y)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    24
by(rule ext)(auto simp: lookup_def update_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    25
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    26
instantiation astate :: (SL_top) SL_top
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    27
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    28
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    29
definition "le_astate F G = (ALL x : set(dom G). lookup F x \<sqsubseteq> fun G x)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    30
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    31
definition
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    32
"join_astate F G =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    33
 FunDom (%x. fun F x \<squnion> fun G x) ([x\<leftarrow>dom F. x : set(dom G)])"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    34
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    35
definition "Top = FunDom (%x. Top) []"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    36
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    37
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    38
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    39
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    40
    apply(auto simp: le_astate_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    41
    by (metis lookup_def preord_class.le_trans top)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    42
qed (auto simp: le_astate_def lookup_def join_astate_def Top_astate_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    43
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    44
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    45
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    46
end