src/HOL/IMP/Collecting.thy
author haftmann
Fri, 01 Nov 2013 18:51:14 +0100
changeset 54230 b1d955791529
parent 52046 bc01725d7918
child 55599 6535c537b243
permissions -rw-r--r--
more simplification rules on unary and binary minus
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
     1
theory Collecting
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
     2
imports Complete_Lattice Big_Step ACom
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
     3
        "~~/src/HOL/ex/Interpretation_with_Defs"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
     4
begin
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
     5
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
     6
subsection "The generic Step function"
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
     7
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
     8
notation
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
     9
  sup (infixl "\<squnion>" 65) and
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    10
  inf (infixl "\<sqinter>" 70) and
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    11
  bot ("\<bottom>") and
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    12
  top ("\<top>")
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    13
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    14
context
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    15
  fixes f :: "vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    16
  fixes g :: "bexp \<Rightarrow> 'a \<Rightarrow> 'a"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    17
begin
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    18
fun Step :: "'a \<Rightarrow> 'a acom \<Rightarrow> 'a acom" where
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    19
"Step S (SKIP {Q}) = (SKIP {S})" |
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    20
"Step S (x ::= e {Q}) =
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    21
  x ::= e {f x e S}" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 52019
diff changeset
    22
"Step S (C1;; C2) = Step S C1;; Step (post C1) C2" |
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    23
"Step S (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) =
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    24
  IF b THEN {g b S} Step P1 C1 ELSE {g (Not b) S} Step P2 C2
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    25
  {post C1 \<squnion> post C2}" |
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    26
"Step S ({I} WHILE b DO {P} C {Q}) =
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    27
  {S \<squnion> post C} WHILE b DO {g b I} Step P C {g (Not b) I}"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    28
end
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    29
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    30
lemma strip_Step[simp]: "strip(Step f g S C) = strip C"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    31
by(induct C arbitrary: S) auto
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    32
8a9f0503b1c0 factored out Step
nipkow
parents: 51039
diff changeset
    33
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    34
subsection "Collecting Semantics of Commands"
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
subsubsection "Annotated commands as a complete lattice"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    37
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    38
instantiation acom :: (order) order
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    39
begin
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    40
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    41
definition less_eq_acom :: "('a::order)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    42
"C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> (\<forall>p<size(annos C1). anno C1 p \<le> anno C2 p)"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    43
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    44
definition less_acom :: "'a acom \<Rightarrow> 'a acom \<Rightarrow> bool" where
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    45
"less_acom x y = (x \<le> y \<and> \<not> y \<le> x)"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    46
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    47
instance
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    48
proof
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    49
  case goal1 show ?case by(simp add: less_acom_def)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    50
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    51
  case goal2 thus ?case by(auto simp: less_eq_acom_def)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    52
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    53
  case goal3 thus ?case by(fastforce simp: less_eq_acom_def size_annos)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    54
next
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    55
  case goal4 thus ?case
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    56
    by(fastforce simp: le_antisym less_eq_acom_def size_annos
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    57
         eq_acom_iff_strip_anno)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    58
qed
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    59
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    60
end
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    61
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    62
lemma less_eq_acom_annos:
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    63
  "C1 \<le> C2 \<longleftrightarrow> strip C1 = strip C2 \<and> list_all2 (op \<le>) (annos C1) (annos C2)"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    64
by(auto simp add: less_eq_acom_def anno_def list_all2_conv_all_nth size_annos_same2)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    65
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    66
lemma SKIP_le[simp]: "SKIP {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = SKIP {S'} \<and> S \<le> S')"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    67
by (cases c) (auto simp:less_eq_acom_def anno_def)
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
    68
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    69
lemma Assign_le[simp]: "x ::= e {S} \<le> c \<longleftrightarrow> (\<exists>S'. c = x ::= e {S'} \<and> S \<le> S')"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    70
by (cases c) (auto simp:less_eq_acom_def anno_def)
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
    71
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 52019
diff changeset
    72
lemma Seq_le[simp]: "C1;;C2 \<le> C \<longleftrightarrow> (\<exists>C1' C2'. C = C1';;C2' \<and> C1 \<le> C1' \<and> C2 \<le> C2')"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    73
apply (cases C)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    74
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    75
done
49095
7df19036392e added annotations after condition in if and while
nipkow
parents: 48759
diff changeset
    76
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    77
lemma If_le[simp]: "IF b THEN {p1} C1 ELSE {p2} C2 {S} \<le> C \<longleftrightarrow>
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    78
  (\<exists>p1' p2' C1' C2' S'. C = IF b THEN {p1'} C1' ELSE {p2'} C2' {S'} \<and>
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    79
     p1 \<le> p1' \<and> p2 \<le> p2' \<and> C1 \<le> C1' \<and> C2 \<le> C2' \<and> S \<le> S')"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    80
apply (cases C)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    81
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    82
done
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
    83
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    84
lemma While_le[simp]: "{I} WHILE b DO {p} C {P} \<le> W \<longleftrightarrow>
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    85
  (\<exists>I' p' C' P'. W = {I'} WHILE b DO {p'} C' {P'} \<and> C \<le> C' \<and> p \<le> p' \<and> I \<le> I' \<and> P \<le> P')"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    86
apply (cases W)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    87
apply(auto simp: less_eq_acom_annos list_all2_append size_annos_same2)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    88
done
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    89
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    90
lemma mono_post: "C \<le> C' \<Longrightarrow> post C \<le> post C'"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    91
using annos_ne[of C']
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    92
by(auto simp: post_def less_eq_acom_def last_conv_nth[OF annos_ne] anno_def
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    93
     dest: size_annos_same)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    94
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    95
definition Inf_acom :: "com \<Rightarrow> 'a::complete_lattice acom set \<Rightarrow> 'a acom" where
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    96
"Inf_acom c M = annotate (\<lambda>p. INF C:M. anno C p) c"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
    97
49397
4f9585401f1a got rid of ad-hoc lift function
nipkow
parents: 49344
diff changeset
    98
interpretation
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
    99
  Complete_Lattice "{C. strip C = c}" "Inf_acom c" for c
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   100
proof
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   101
  case goal1 thus ?case
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   102
    by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_lower)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   103
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   104
  case goal2 thus ?case
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   105
    by(auto simp: Inf_acom_def less_eq_acom_def size_annos intro:INF_greatest)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   106
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   107
  case goal3 thus ?case by(auto simp: Inf_acom_def)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   108
qed
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   109
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   110
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   111
subsubsection "Collecting semantics"
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   112
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   113
definition "step = Step (\<lambda>x e S. {s(x := aval e s) |s. s : S}) (\<lambda>b S. {s:S. bval b s})"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   114
46070
nipkow
parents: 46069
diff changeset
   115
definition CS :: "com \<Rightarrow> state set acom" where
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents: 47818
diff changeset
   116
"CS c = lfp c (step UNIV)"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   117
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   118
lemma mono2_Step: fixes C1 C2 :: "'a::semilattice_sup acom"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   119
  assumes "!!x e S1 S2. S1 \<le> S2 \<Longrightarrow> f x e S1 \<le> f x e S2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   120
          "!!b S1 S2. S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   121
  shows "C1 \<le> C2 \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   122
proof(induction S1 C1 arbitrary: C2 S2 rule: Step.induct)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   123
  case 1 thus ?case by(auto)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   124
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   125
  case 2 thus ?case by (auto simp: assms(1))
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   126
next
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   127
  case 3 thus ?case by(auto simp: mono_post)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   128
next
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   129
  case 4 thus ?case
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   130
    by(auto simp: subset_iff assms(2))
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   131
      (metis mono_post le_supI1 le_supI2)+
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   132
next
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   133
  case 5 thus ?case
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   134
    by(auto simp: subset_iff assms(2))
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   135
      (metis mono_post le_supI1 le_supI2)+
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   136
qed
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   137
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   138
lemma mono2_step: "C1 \<le> C2 \<Longrightarrow> S1 \<subseteq> S2 \<Longrightarrow> step S1 C1 \<le> step S2 C2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   139
unfolding step_def by(rule mono2_Step) auto
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   140
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   141
lemma mono_step: "mono (step S)"
46334
nipkow
parents: 46116
diff changeset
   142
by(blast intro: monoI mono2_step)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   143
49095
7df19036392e added annotations after condition in if and while
nipkow
parents: 48759
diff changeset
   144
lemma strip_step: "strip(step S C) = strip C"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   145
by (induction C arbitrary: S) (auto simp: step_def)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   146
48759
ff570720ba1c Improved complete lattice formalisation - no more index set.
nipkow
parents: 47818
diff changeset
   147
lemma lfp_cs_unfold: "lfp c (step S) = step S (lfp c (step S))"
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
   148
apply(rule lfp_unfold[OF _  mono_step])
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
   149
apply(simp add: strip_step)
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
   150
done
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   151
46070
nipkow
parents: 46069
diff changeset
   152
lemma CS_unfold: "CS c = step UNIV (CS c)"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   153
by (metis CS_def lfp_cs_unfold)
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   154
46070
nipkow
parents: 46069
diff changeset
   155
lemma strip_CS[simp]: "strip(CS c) = c"
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45885
diff changeset
   156
by(simp add: CS_def index_lfp[simplified])
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   157
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   158
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   159
subsubsection "Relation to big-step semantics"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   160
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   161
lemma asize_nz: "asize(c::com) \<noteq> 0"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   162
by (metis length_0_conv length_annos_annotate annos_ne)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   163
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   164
lemma post_Inf_acom:
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   165
  "\<forall>C\<in>M. strip C = c \<Longrightarrow> post (Inf_acom c M) = \<Inter>(post ` M)"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   166
apply(subgoal_tac "\<forall>C\<in>M. size(annos C) = asize c")
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   167
 apply(simp add: post_anno_asize Inf_acom_def asize_nz neq0_conv[symmetric])
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   168
apply(simp add: size_annos)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   169
done
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   170
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   171
lemma post_lfp: "post(lfp c f) = (\<Inter>{post C|C. strip C = c \<and> f C \<le> C})"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   172
by(auto simp add: lfp_def post_Inf_acom)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   173
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   174
lemma big_step_post_step:
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   175
  "\<lbrakk> (c, s) \<Rightarrow> t;  strip C = c;  s \<in> S;  step S C \<le> C \<rbrakk> \<Longrightarrow> t \<in> post C"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   176
proof(induction arbitrary: C S rule: big_step_induct)
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   177
  case Skip thus ?case by(auto simp: strip_eq_SKIP step_def post_def)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   178
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   179
  case Assign thus ?case
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   180
    by(fastforce simp: strip_eq_Assign step_def post_def)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   181
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   182
  case Seq thus ?case
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   183
    by(fastforce simp: strip_eq_Seq step_def post_def last_append annos_ne)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   184
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   185
  case IfTrue thus ?case apply(auto simp: strip_eq_If step_def post_def)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   186
    by (metis (lifting,full_types) mem_Collect_eq set_mp)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   187
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   188
  case IfFalse thus ?case apply(auto simp: strip_eq_If step_def post_def)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   189
    by (metis (lifting,full_types) mem_Collect_eq set_mp)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   190
next
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   191
  case (WhileTrue b s1 c' s2 s3)
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   192
  from WhileTrue.prems(1) obtain I P C' Q where "C = {I} WHILE b DO {P} C' {Q}" "strip C' = c'"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   193
    by(auto simp: strip_eq_While)
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   194
  from WhileTrue.prems(3) `C = _`
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   195
  have "step P C' \<le> C'"  "{s \<in> I. bval b s} \<le> P"  "S \<le> I"  "step (post C') C \<le> C"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   196
    by (auto simp: step_def post_def)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   197
  have "step {s \<in> I. bval b s} C' \<le> C'"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   198
    by (rule order_trans[OF mono2_step[OF order_refl `{s \<in> I. bval b s} \<le> P`] `step P C' \<le> C'`])
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   199
  have "s1: {s:I. bval b s}" using `s1 \<in> S` `S \<subseteq> I` `bval b s1` by auto
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   200
  note s2_in_post_C' = WhileTrue.IH(1)[OF `strip C' = c'` this `step {s \<in> I. bval b s} C' \<le> C'`]
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   201
  from WhileTrue.IH(2)[OF WhileTrue.prems(1) s2_in_post_C' `step (post C') C \<le> C`]
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   202
  show ?case .
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   203
next
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   204
  case (WhileFalse b s1 c') thus ?case
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51630
diff changeset
   205
    by (force simp: strip_eq_While step_def post_def)
49487
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   206
qed
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   207
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   208
lemma big_step_lfp: "\<lbrakk> (c,s) \<Rightarrow> t;  s \<in> S \<rbrakk> \<Longrightarrow> t \<in> post(lfp c (step S))"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   209
by(auto simp add: post_lfp intro: big_step_post_step)
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   210
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   211
lemma big_step_CS: "(c,s) \<Rightarrow> t \<Longrightarrow> t : post(CS c)"
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   212
by(simp add: CS_def big_step_lfp)
7e7ac4956117 conected CS to big-step
nipkow
parents: 49397
diff changeset
   213
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents:
diff changeset
   214
end