src/HOL/IMP/Abs_Int0.thy
author nipkow
Wed, 18 Jan 2012 10:05:14 +0100
changeset 46246 e69684c1c142
parent 46158 8b5f1f91ef38
child 46334 3858dc8eabd8
permissions -rw-r--r--
introduced commands over a set of vars
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     2
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     3
theory Abs_Int0
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     4
imports Abs_State
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     5
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     6
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     7
subsection "Computable Abstract Interpretation"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
     8
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
     9
text{* Abstract interpretation over type @{text st} instead of
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    10
functions. *}
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    11
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    12
context Val_abs
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    13
begin
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    14
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    15
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
46039
nipkow
parents: 45903
diff changeset
    16
"aval' (N n) S = num' n" |
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    17
"aval' (V x) S = lookup S x" |
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    18
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    19
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    20
lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
46039
nipkow
parents: 45903
diff changeset
    21
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def lookup_def)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    22
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    23
end
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    24
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    25
text{* The for-clause (here and elsewhere) only serves the purpose of fixing
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    26
the name of the type parameter @{typ 'av} which would otherwise be renamed to
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    27
@{typ 'a}. *}
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    28
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    29
locale Abs_Int = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    30
begin
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    31
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    32
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    33
"step' S (SKIP {P}) = (SKIP {S})" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    34
"step' S (x ::= e {P}) =
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    35
  x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S))}" |
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    36
"step' S (c1; c2) = step' S c1; step' (post c1) c2" |
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    37
"step' S (IF b THEN c1 ELSE c2 {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    38
  (let c1' = step' S c1; c2' = step' S c2
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    39
   in IF b THEN c1' ELSE c2' {post c1 \<squnion> post c2})" |
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    40
"step' S ({Inv} WHILE b DO c {P}) =
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    41
   {S \<squnion> post c} WHILE b DO step' Inv c {Inv}"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    42
46063
81ebd0cdb300 tuned types
nipkow
parents: 46039
diff changeset
    43
definition AI :: "com \<Rightarrow> 'av st option acom option" where
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    44
"AI = lpfp\<^isub>c (step' \<top>)"
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
    45
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
    46
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    47
lemma strip_step'[simp]: "strip(step' S c) = strip c"
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    48
by(induct c arbitrary: S) (simp_all add: Let_def)
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    49
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    50
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
    51
text{* Soundness: *}
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    52
46039
nipkow
parents: 45903
diff changeset
    53
lemma in_gamma_update:
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    54
  "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
46039
nipkow
parents: 45903
diff changeset
    55
by(simp add: \<gamma>_st_def lookup_update)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    56
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    57
text{* The soundness proofs are textually identical to the ones for the step
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    58
function operating on states as functions. *}
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    59
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    60
lemma step_preserves_le:
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    61
  "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; cs \<le> \<gamma>\<^isub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' S' ca)"
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    62
proof(induction cs arbitrary: ca S S')
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    63
  case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    64
next
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    65
  case Assign thus ?case
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    66
    by (fastforce simp: Assign_le  map_acom_Assign intro: aval'_sound in_gamma_update
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    67
      split: option.splits del:subsetD)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    68
next
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    69
  case Semi thus ?case apply (auto simp: Semi_le map_acom_Semi)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    70
    by (metis le_post post_map_acom)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    71
next
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    72
  case (If b cs1 cs2 P)
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    73
  then obtain ca1 ca2 Pa where
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    74
      "ca= IF b THEN ca1 ELSE ca2 {Pa}"
46039
nipkow
parents: 45903
diff changeset
    75
      "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    76
    by (fastforce simp: If_le map_acom_If)
46039
nipkow
parents: 45903
diff changeset
    77
  moreover have "post cs1 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
nipkow
parents: 45903
diff changeset
    78
    by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
nipkow
parents: 45903
diff changeset
    79
  moreover have "post cs2 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
nipkow
parents: 45903
diff changeset
    80
    by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    81
  ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    82
next
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    83
  case (While I b cs1 P)
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    84
  then obtain ca1 Ia Pa where
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    85
    "ca = {Ia} WHILE b DO ca1 {Pa}"
46039
nipkow
parents: 45903
diff changeset
    86
    "I \<subseteq> \<gamma>\<^isub>o Ia" "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
    87
    by (fastforce simp: map_acom_While While_le)
46067
a03bf644cb27 tuned var names
nipkow
parents: 46066
diff changeset
    88
  moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post ca1)"
a03bf644cb27 tuned var names
nipkow
parents: 46066
diff changeset
    89
    using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified]
46039
nipkow
parents: 45903
diff changeset
    90
    by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    91
  ultimately show ?case by (simp add: While.IH subset_iff)
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    92
qed
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
    93
46070
nipkow
parents: 46068
diff changeset
    94
lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    95
proof(simp add: CS_def AI_def)
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    96
  assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    97
  have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
    98
  have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
    99
    by(simp add: strip_lpfpc[OF _ 1])
46066
e81411bfa7ef tuned argument order
nipkow
parents: 46063
diff changeset
   100
  have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
45903
02dd9319dcb7 improved indexed complete lattice
nipkow
parents: 45655
diff changeset
   101
  proof(rule lfp_lowerbound[simplified,OF 3])
45655
a49f9428aba4 simplified Collecting1 and renamed: step -> step', step_cs -> step
nipkow
parents: 45623
diff changeset
   102
    show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
46068
b9d4ec0f79ac tuned proofs
nipkow
parents: 46067
diff changeset
   103
    proof(rule step_preserves_le[OF _ _])
46039
nipkow
parents: 45903
diff changeset
   104
      show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
nipkow
parents: 45903
diff changeset
   105
      show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
   106
    qed
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
   107
  qed
46066
e81411bfa7ef tuned argument order
nipkow
parents: 46063
diff changeset
   108
  from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
46039
nipkow
parents: 45903
diff changeset
   109
    by (blast intro: mono_gamma_c order_trans)
45623
f682f3f7b726 Abstract interpretation is now based uniformly on annotated programs,
nipkow
parents: 45127
diff changeset
   110
qed
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   112
end
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   113
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   114
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   115
subsubsection "Monotonicity"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   116
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   117
locale Abs_Int_mono = Abs_Int +
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   118
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   119
begin
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   120
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   121
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   122
by(induction e) (auto simp: le_st_def lookup_def mono_plus')
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   123
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   124
lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> update S x a \<sqsubseteq> update S' x a'"
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   125
by(auto simp add: le_st_def lookup_def update_def)
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   126
46153
nipkow
parents: 46070
diff changeset
   127
lemma mono_step': "S \<sqsubseteq> S' \<Longrightarrow> c \<sqsubseteq> c' \<Longrightarrow> step' S c \<sqsubseteq> step' S' c'"
nipkow
parents: 46070
diff changeset
   128
apply(induction c c' arbitrary: S S' rule: le_acom.induct)
nipkow
parents: 46070
diff changeset
   129
apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj
nipkow
parents: 46070
diff changeset
   130
            split: option.split)
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   131
done
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   132
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents:
diff changeset
   133
end
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   134
46158
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   135
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   136
subsubsection "Ascending Chain Condition"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   137
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   138
hide_const acc
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   139
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   140
abbreviation "strict r == r \<inter> -(r^-1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   141
abbreviation "acc r == wf((strict r)^-1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   142
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   143
lemma strict_inv_image: "strict(inv_image r f) = inv_image (strict r) f"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   144
by(auto simp: inv_image_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   145
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   146
lemma acc_inv_image:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   147
  "acc r \<Longrightarrow> acc (inv_image r f)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   148
by (metis converse_inv_image strict_inv_image wf_inv_image)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   149
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   150
text{* ACC for option type: *}
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   151
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   152
lemma acc_option: assumes "acc {(x,y::'a::preord). x \<sqsubseteq> y}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   153
shows "acc {(x,y::'a option). x \<sqsubseteq> y}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   154
proof(auto simp: wf_eq_minimal)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   155
  fix xo :: "'a option" and Qo assume "xo : Qo"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   156
  let ?Q = "{x. Some x \<in> Qo}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   157
  show "\<exists>yo\<in>Qo. \<forall>zo. yo \<sqsubseteq> zo \<and> ~ zo \<sqsubseteq> yo \<longrightarrow> zo \<notin> Qo" (is "\<exists>zo\<in>Qo. ?P zo")
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   158
  proof cases
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   159
    assume "?Q = {}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   160
    hence "?P None" by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   161
    moreover have "None \<in> Qo" using `?Q = {}` `xo : Qo`
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   162
      by auto (metis not_Some_eq)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   163
    ultimately show ?thesis by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   164
  next
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   165
    assume "?Q \<noteq> {}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   166
    with assms show ?thesis
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   167
      apply(auto simp: wf_eq_minimal)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   168
      apply(erule_tac x="?Q" in allE)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   169
      apply auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   170
      apply(rule_tac x = "Some z" in bexI)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   171
      by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   172
  qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   173
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   174
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   175
text{* ACC for abstract states, via measure functions. *}
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   176
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   177
(*FIXME mv*)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   178
lemma setsum_strict_mono1:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   179
fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   180
assumes "finite A" and "ALL x:A. f x \<le> g x" and "EX a:A. f a < g a"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   181
shows "setsum f A < setsum g A"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   182
proof-
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   183
  from assms(3) obtain a where a: "a:A" "f a < g a" by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   184
  have "setsum f A = setsum f ((A-{a}) \<union> {a})"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   185
    by(simp add:insert_absorb[OF `a:A`])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   186
  also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   187
    using `finite A` by(subst setsum_Un_disjoint) auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   188
  also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   189
    by(rule setsum_mono)(simp add: assms(2))
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   190
  also have "setsum f {a} < setsum g {a}" using a by simp
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   191
  also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   192
    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   193
  also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   194
  finally show ?thesis by (metis add_right_mono add_strict_left_mono)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   195
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   196
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   197
lemma measure_st: assumes "(strict{(x,y::'a::SL_top). x \<sqsubseteq> y})^-1 <= measure m"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   198
and "\<forall>x y::'a. x \<sqsubseteq> y \<and> y \<sqsubseteq> x \<longrightarrow> m x = m y"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   199
shows "(strict{(S,S'::'a st). S \<sqsubseteq> S'})^-1 \<subseteq>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   200
  measure(%fd. \<Sum>x| x\<in>set(dom fd) \<and> ~ \<top> \<sqsubseteq> fun fd x. m(fun fd x)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   201
proof-
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   202
  { fix S S' :: "'a st" assume "S \<sqsubseteq> S'" "~ S' \<sqsubseteq> S"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   203
    let ?X = "set(dom S)" let ?Y = "set(dom S')"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   204
    let ?f = "fun S" let ?g = "fun S'"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   205
    let ?X' = "{x:?X. ~ \<top> \<sqsubseteq> ?f x}" let ?Y' = "{y:?Y. ~ \<top> \<sqsubseteq> ?g y}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   206
    from `S \<sqsubseteq> S'` have "ALL y:?Y'\<inter>?X. ?f y \<sqsubseteq> ?g y"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   207
      by(auto simp: le_st_def lookup_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   208
    hence 1: "ALL y:?Y'\<inter>?X. m(?g y)+1 \<le> m(?f y)+1"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   209
      using assms(1,2) by(fastforce)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   210
    from `~ S' \<sqsubseteq> S` obtain u where u: "u : ?X" "~ lookup S' u \<sqsubseteq> ?f u"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   211
      by(auto simp: le_st_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   212
    hence "u : ?X'" by simp (metis preord_class.le_trans top)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   213
    have "?Y'-?X = {}" using `S \<sqsubseteq> S'` by(fastforce simp: le_st_def lookup_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   214
    have "?Y'\<inter>?X <= ?X'" apply auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   215
      apply (metis `S \<sqsubseteq> S'` le_st_def lookup_def preord_class.le_trans)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   216
      done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   217
    have "(\<Sum>y\<in>?Y'. m(?g y)+1) = (\<Sum>y\<in>(?Y'-?X) \<union> (?Y'\<inter>?X). m(?g y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   218
      by (metis Un_Diff_Int)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   219
    also have "\<dots> = (\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   220
      using `?Y'-?X = {}` by (metis Un_empty_left)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   221
    also have "\<dots> < (\<Sum>x\<in>?X'. m(?f x)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   222
    proof cases
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   223
      assume "u \<in> ?Y'"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   224
      hence "m(?g u) < m(?f u)" using assms(1) `S \<sqsubseteq> S'` u
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   225
        by (fastforce simp: le_st_def lookup_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   226
      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) < (\<Sum>y\<in>?Y'\<inter>?X. m(?f y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   227
        using `u:?X` `u:?Y'` `m(?g u) < m(?f u)`
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   228
        by(fastforce intro!: setsum_strict_mono1[OF _ 1])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   229
      also have "\<dots> \<le> (\<Sum>y\<in>?X'. m(?f y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   230
        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   231
      finally show ?thesis .
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   232
    next
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   233
      assume "u \<notin> ?Y'"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   234
      with `?Y'\<inter>?X <= ?X'` have "?Y'\<inter>?X - {u} <= ?X' - {u}" by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   235
      have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) = (\<Sum>y\<in>?Y'\<inter>?X - {u}. m(?g y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   236
      proof-
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   237
        have "?Y'\<inter>?X = ?Y'\<inter>?X - {u}" using `u \<notin> ?Y'` by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   238
        thus ?thesis by metis
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   239
      qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   240
      also have "\<dots> < (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) + (\<Sum>y\<in>{u}. m(?f y)+1)" by simp
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   241
      also have "(\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) \<le> (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?f y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   242
        using 1 by(blast intro: setsum_mono)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   243
      also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   244
        by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   245
      also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   246
        using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   247
      also have "\<dots> = (\<Sum>x\<in>?X'. m(?f x)+1)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   248
        using `u : ?X'` by(simp add:insert_absorb)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   249
      finally show ?thesis by (blast intro: add_right_mono)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   250
    qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   251
    finally have "(\<Sum>y\<in>?Y'. m(?g y)+1) < (\<Sum>x\<in>?X'. m(?f x)+1)" .
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   252
  } thus ?thesis by(auto simp add: measure_def inv_image_def)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   253
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   254
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   255
text{* ACC for acom. First the ordering on acom is related to an ordering on
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   256
lists of annotations. *}
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   257
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   258
(* FIXME mv and add [simp] *)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   259
lemma listrel_Cons_iff:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   260
  "(x#xs, y#ys) : listrel r \<longleftrightarrow> (x,y) \<in> r \<and> (xs,ys) \<in> listrel r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   261
by (blast intro:listrel.Cons)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   262
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   263
lemma listrel_app: "(xs1,ys1) : listrel r \<Longrightarrow> (xs2,ys2) : listrel r
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   264
  \<Longrightarrow> (xs1@xs2, ys1@ys2) : listrel r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   265
by(auto simp add: listrel_iff_zip)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   266
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   267
lemma listrel_app_same_size: "size xs1 = size ys1 \<Longrightarrow> size xs2 = size ys2 \<Longrightarrow>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   268
  (xs1@xs2, ys1@ys2) : listrel r \<longleftrightarrow>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   269
  (xs1,ys1) : listrel r \<and> (xs2,ys2) : listrel r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   270
by(auto simp add: listrel_iff_zip)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   271
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   272
lemma listrel_converse: "listrel(r^-1) = (listrel r)^-1"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   273
proof-
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   274
  { fix xs ys
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   275
    have "(xs,ys) : listrel(r^-1) \<longleftrightarrow> (ys,xs) : listrel r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   276
      apply(induct xs arbitrary: ys)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   277
      apply (fastforce simp: listrel.Nil)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   278
      apply (fastforce simp: listrel_Cons_iff)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   279
      done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   280
  } thus ?thesis by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   281
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   282
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   283
(* It would be nice to get rid of refl & trans and build them into the proof *)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   284
lemma acc_listrel: fixes r :: "('a*'a)set" assumes "refl r" and "trans r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   285
and "acc r" shows "acc (listrel r - {([],[])})"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   286
proof-
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   287
  have refl: "!!x. (x,x) : r" using `refl r` unfolding refl_on_def by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   288
  have trans: "!!x y z. (x,y) : r \<Longrightarrow> (y,z) : r \<Longrightarrow> (x,z) : r"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   289
    using `trans r` unfolding trans_def by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   290
  from assms(3) obtain mx :: "'a set \<Rightarrow> 'a" where
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   291
    mx: "!!S x. x:S \<Longrightarrow> mx S : S \<and> (\<forall>y. (mx S,y) : strict r \<longrightarrow> y \<notin> S)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   292
    by(simp add: wf_eq_minimal) metis
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   293
  let ?R = "listrel r - {([], [])}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   294
  { fix Q and xs :: "'a list"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   295
    have "xs \<in> Q \<Longrightarrow> \<exists>ys. ys\<in>Q \<and> (\<forall>zs. (ys, zs) \<in> strict ?R \<longrightarrow> zs \<notin> Q)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   296
      (is "_ \<Longrightarrow> \<exists>ys. ?P Q ys")
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   297
    proof(induction xs arbitrary: Q rule: length_induct)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   298
      case (1 xs)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   299
      { have "!!ys Q. size ys < size xs \<Longrightarrow> ys : Q \<Longrightarrow> EX ms. ?P Q ms"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   300
          using "1.IH" by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   301
      } note IH = this
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   302
      show ?case
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   303
      proof(cases xs)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   304
        case Nil with `xs : Q` have "?P Q []" by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   305
        thus ?thesis by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   306
      next
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   307
        case (Cons x ys)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   308
        let ?Q1 = "{a. \<exists>bs. size bs = size ys \<and> a#bs : Q}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   309
        have "x : ?Q1" using `xs : Q` Cons by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   310
        from mx[OF this] obtain m1 where
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   311
          1: "m1 \<in> ?Q1 \<and> (\<forall>y. (m1,y) \<in> strict r \<longrightarrow> y \<notin> ?Q1)" by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   312
        then obtain ms1 where "size ms1 = size ys" "m1#ms1 : Q" by blast+
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   313
        hence "size ms1 < size xs" using Cons by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   314
        let ?Q2 = "{bs. \<exists>m1'. (m1',m1):r \<and> (m1,m1'):r \<and> m1'#bs : Q \<and> size bs = size ms1}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   315
        have "ms1 : ?Q2" using `m1#ms1 : Q` by(blast intro: refl)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   316
        from IH[OF `size ms1 < size xs` this]
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   317
        obtain ms where 2: "?P ?Q2 ms" by auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   318
        then obtain m1' where m1': "(m1',m1) : r \<and> (m1,m1') : r \<and> m1'#ms : Q"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   319
          by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   320
        hence "\<forall>ab. (m1'#ms,ab) : strict ?R \<longrightarrow> ab \<notin> Q" using 1 2
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   321
          apply (auto simp: listrel_Cons_iff)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   322
          apply (metis `length ms1 = length ys` listrel_eq_len trans)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   323
          by (metis `length ms1 = length ys` listrel_eq_len trans)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   324
        with m1' show ?thesis by blast
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   325
      qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   326
    qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   327
  }
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   328
  thus ?thesis unfolding wf_eq_minimal by (metis converse_iff)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   329
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   330
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   331
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   332
fun annos :: "'a acom \<Rightarrow> 'a list" where
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   333
"annos (SKIP {a}) = [a]" |
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   334
"annos (x::=e {a}) = [a]" |
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   335
"annos (c1;c2) = annos c1 @ annos c2" |
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   336
"annos (IF b THEN c1 ELSE c2 {a}) = a #  annos c1 @ annos c2" |
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   337
"annos ({i} WHILE b DO c {a}) = i # a # annos c"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   338
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   339
lemma size_annos_same: "strip c1 = strip c2 \<Longrightarrow> size(annos c1) = size(annos c2)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   340
apply(induct c2 arbitrary: c1)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   341
apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Semi strip_eq_If strip_eq_While)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   342
done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   343
46246
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   344
lemmas size_annos_same2 = eqTrueI[OF size_annos_same]
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   345
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   346
lemma set_annos_anno: "set (annos (anno a c)) = {a}"
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   347
by(induction c)(auto)
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   348
46158
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   349
lemma le_iff_le_annos: "c1 \<sqsubseteq> c2 \<longleftrightarrow>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   350
 (annos c1, annos c2) : listrel{(x,y). x \<sqsubseteq> y} \<and> strip c1 = strip c2"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   351
apply(induct c1 c2 rule: le_acom.induct)
46246
e69684c1c142 introduced commands over a set of vars
nipkow
parents: 46158
diff changeset
   352
apply (auto simp: listrel.Nil listrel_Cons_iff listrel_app size_annos_same2)
46158
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   353
apply (metis listrel_app_same_size size_annos_same)+
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   354
done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   355
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   356
lemma le_acom_subset_same_annos:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   357
 "(strict{(c,c'::'a::preord acom). c \<sqsubseteq> c'})^-1 \<subseteq>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   358
  (strict(inv_image (listrel{(a,a'::'a). a \<sqsubseteq> a'} - {([],[])}) annos))^-1"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   359
by(auto simp: le_iff_le_annos)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   360
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   361
lemma acc_acom: "acc {(a,a'::'a::preord). a \<sqsubseteq> a'} \<Longrightarrow>
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   362
  acc {(c,c'::'a acom). c \<sqsubseteq> c'}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   363
apply(rule wf_subset[OF _ le_acom_subset_same_annos])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   364
apply(rule acc_inv_image[OF acc_listrel])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   365
apply(auto simp: refl_on_def trans_def intro: le_trans)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   366
done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   367
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   368
text{* Termination of the fixed-point finders, assuming monotone functions: *}
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   369
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   370
lemma pfp_termination:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   371
fixes x0 :: "'a::preord"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   372
assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" and "acc {(x::'a,y). x \<sqsubseteq> y}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   373
and "x0 \<sqsubseteq> f x0" shows "EX x. pfp f x0 = Some x"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   374
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. x \<sqsubseteq> f x"])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   375
  show "wf {(x, s). (s \<sqsubseteq> f s \<and> \<not> f s \<sqsubseteq> s) \<and> x = f s}"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   376
    by(rule wf_subset[OF assms(2)]) auto
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   377
next
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   378
  show "x0 \<sqsubseteq> f x0" by(rule assms)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   379
next
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   380
  fix x assume "x \<sqsubseteq> f x" thus "f x \<sqsubseteq> f(f x)" by(rule mono)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   381
qed
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   382
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   383
lemma lpfpc_termination:
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   384
  fixes f :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom)"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   385
  assumes "acc {(x::'a,y). x \<sqsubseteq> y}" and "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   386
  and "\<And>c. strip(f c) = strip c"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   387
  shows "\<exists>c'. lpfp\<^isub>c f c = Some c'"
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   388
unfolding lpfp\<^isub>c_def
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   389
apply(rule pfp_termination)
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   390
  apply(erule assms(2))
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   391
 apply(rule acc_acom[OF acc_option[OF assms(1)]])
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   392
apply(simp add: bot_acom assms(3))
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   393
done
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   394
8b5f1f91ef38 Added termination to IMP Abs_Int
nipkow
parents: 46153
diff changeset
   395
45127
d2eb07a1e01b separated monotonicity reasoning and defined narrowing with while_option
nipkow
parents: 45111
diff changeset
   396
end