src/HOL/IMP/Abs_Int0.thy
author wenzelm
Wed, 26 Dec 2018 16:25:20 +0100
changeset 69505 cc2d676d5395
parent 68778 4566bac4517d
child 69597 ff784d5a5bfb
permissions -rw-r--r--
isabelle update_cartouches -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     2
68778
4566bac4517d improved sectioning
nipkow
parents: 68484
diff changeset
     3
subsection "Abstract Interpretation"
4566bac4517d improved sectioning
nipkow
parents: 68484
diff changeset
     4
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     5
theory Abs_Int0
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     6
imports Abs_Int_init
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     7
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     8
68778
4566bac4517d improved sectioning
nipkow
parents: 68484
diff changeset
     9
subsubsection "Orderings"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    10
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
    11
text\<open>The basic type classes @{class order}, @{class semilattice_sup} and @{class order_top} are
68484
59793df7f853 clarified document antiquotation @{theory};
wenzelm
parents: 67613
diff changeset
    12
defined in @{theory Main}, more precisely in theories @{theory HOL.Orderings} and @{theory HOL.Lattices}.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
    13
If you view this theory with jedit, just click on the names to get there.\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    14
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52504
diff changeset
    15
class semilattice_sup_top = semilattice_sup + order_top
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    16
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    17
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
    18
instance "fun" :: (type, semilattice_sup_top) semilattice_sup_top ..
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    19
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    20
instantiation option :: (order)order
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    21
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    22
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    23
fun less_eq_option where
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    24
"Some x \<le> Some y = (x \<le> y)" |
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    25
"None \<le> y = True" |
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    26
"Some _ \<le> None = False"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    27
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    28
definition less_option where "x < (y::'a option) = (x \<le> y \<and> \<not> y \<le> x)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    29
51628
nipkow
parents: 51625
diff changeset
    30
lemma le_None[simp]: "(x \<le> None) = (x = None)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    31
by (cases x) simp_all
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    32
51628
nipkow
parents: 51625
diff changeset
    33
lemma Some_le[simp]: "(Some x \<le> u) = (\<exists>y. u = Some y \<and> x \<le> y)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    34
by (cases u) auto
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    35
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    36
instance
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    37
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    38
  case 1 show ?case by(rule less_option_def)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    39
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    40
  case (2 x) show ?case by(cases x, simp_all)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    41
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    42
  case (3 x y z) thus ?case by(cases z, simp, cases y, simp, cases x, auto)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    43
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    44
  case (4 x y) thus ?case by(cases y, simp, cases x, auto)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    45
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    46
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    47
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    48
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51359
diff changeset
    49
instantiation option :: (sup)sup
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    50
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    51
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51359
diff changeset
    52
fun sup_option where
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    53
"Some x \<squnion> Some y = Some(x \<squnion> y)" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    54
"None \<squnion> y = y" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    55
"x \<squnion> None = x"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    56
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51359
diff changeset
    57
lemma sup_None2[simp]: "x \<squnion> None = x"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    58
by (cases x) simp_all
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    59
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    60
instance ..
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    61
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    62
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    63
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
    64
instantiation option :: (semilattice_sup_top)semilattice_sup_top
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    65
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    66
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    67
definition top_option where "\<top> = Some \<top>"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    68
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    69
instance
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    70
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    71
  case (4 a) show ?case by(cases a, simp_all add: top_option_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    72
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    73
  case (1 x y) thus ?case by(cases x, simp, cases y, simp_all)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    74
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    75
  case (2 x y) thus ?case by(cases y, simp, cases x, simp_all)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    76
next
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    77
  case (3 x y z) thus ?case by(cases z, simp, cases y, simp, cases x, simp_all)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    78
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    79
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    80
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    81
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    82
lemma [simp]: "(Some x < Some y) = (x < y)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    83
by(auto simp: less_le)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    84
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52504
diff changeset
    85
instantiation option :: (order)order_bot
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    86
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    87
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49344
diff changeset
    88
definition bot_option :: "'a option" where
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    89
"\<bottom> = None"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    90
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    91
instance
61179
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    92
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 60974
diff changeset
    93
  case 1 thus ?case by(auto simp: bot_option_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    94
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    95
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    96
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    97
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    98
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    99
definition bot :: "com \<Rightarrow> 'a option acom" where
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   100
"bot c = annotate (\<lambda>p. None) c"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   101
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   102
lemma bot_least: "strip C = c \<Longrightarrow> bot c \<le> C"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   103
by(auto simp: bot_def less_eq_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   104
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   105
lemma strip_bot[simp]: "strip(bot c) = c"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   106
by(simp add: bot_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   107
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   108
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   109
subsubsection "Pre-fixpoint iteration"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   110
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   111
definition pfp :: "(('a::order) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   112
"pfp f = while_option (\<lambda>x. \<not> f x \<le> x) f"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   113
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   114
lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<le> x"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   115
using while_option_stop[OF assms[simplified pfp_def]] by simp
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   116
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   117
lemma while_least:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   118
fixes q :: "'a::order"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   119
assumes "\<forall>x\<in>L.\<forall>y\<in>L. x \<le> y \<longrightarrow> f x \<le> f y" and "\<forall>x. x \<in> L \<longrightarrow> f x \<in> L"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   120
and "\<forall>x \<in> L. b \<le> x" and "b \<in> L" and "f q \<le> q" and "q \<in> L"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   121
and "while_option P f b = Some p"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   122
shows "p \<le> q"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   123
using while_option_rule[OF _  assms(7)[unfolded pfp_def],
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   124
                        where P = "%x. x \<in> L \<and> x \<le> q"]
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   125
by (metis assms(1-6) order_trans)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   126
51710
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   127
lemma pfp_bot_least:
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   128
assumes "\<forall>x\<in>{C. strip C = c}.\<forall>y\<in>{C. strip C = c}. x \<le> y \<longrightarrow> f x \<le> f y"
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   129
and "\<forall>C. C \<in> {C. strip C = c} \<longrightarrow> f C \<in> {C. strip C = c}"
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   130
and "f C' \<le> C'" "strip C' = c" "pfp f (bot c) = Some C"
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   131
shows "C \<le> C'"
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   132
by(rule while_least[OF assms(1,2) _ _ assms(3) _ assms(5)[unfolded pfp_def]])
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   133
  (simp_all add: assms(4) bot_least)
f692657e0f71 moved leastness lemma
nipkow
parents: 51694
diff changeset
   134
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   135
lemma pfp_inv:
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   136
  "pfp f x = Some y \<Longrightarrow> (\<And>x. P x \<Longrightarrow> P(f x)) \<Longrightarrow> P x \<Longrightarrow> P y"
58955
1694bad18568 avoid erule and rotated in IMP
nipkow
parents: 58947
diff changeset
   137
unfolding pfp_def by (blast intro: while_option_rule)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   138
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   139
lemma strip_pfp:
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   140
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   141
using pfp_inv[OF assms(2), where P = "%x. g x = g x0"] assms(1) by simp
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   142
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   143
68778
4566bac4517d improved sectioning
nipkow
parents: 68484
diff changeset
   144
subsubsection "Abstract Interpretation"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   145
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   146
definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   147
"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   148
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   149
fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   150
"\<gamma>_option \<gamma> None = {}" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   151
"\<gamma>_option \<gamma> (Some a) = \<gamma> a"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   152
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   153
text\<open>The interface for abstract values:\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   154
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52046
diff changeset
   155
locale Val_semilattice =
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   156
fixes \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   157
  assumes mono_gamma: "a \<le> b \<Longrightarrow> \<gamma> a \<le> \<gamma> b"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   158
  and gamma_Top[simp]: "\<gamma> \<top> = UNIV"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   159
fixes num' :: "val \<Rightarrow> 'av"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   160
and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av"
51036
e7b54119c436 tuned top
nipkow
parents: 50986
diff changeset
   161
  assumes gamma_num': "i \<in> \<gamma>(num' i)"
e7b54119c436 tuned top
nipkow
parents: 50986
diff changeset
   162
  and gamma_plus': "i1 \<in> \<gamma> a1 \<Longrightarrow> i2 \<in> \<gamma> a2 \<Longrightarrow> i1+i2 \<in> \<gamma>(plus' a1 a2)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   163
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   164
type_synonym 'av st = "(vname \<Rightarrow> 'av)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   165
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   166
text\<open>The for-clause (here and elsewhere) only serves the purpose of fixing
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   167
the name of the type parameter @{typ 'av} which would otherwise be renamed to
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   168
@{typ 'a}.\<close>
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   169
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52046
diff changeset
   170
locale Abs_Int_fun = Val_semilattice where \<gamma>=\<gamma>
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   171
  for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   172
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   173
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   174
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
50896
nipkow
parents: 49497
diff changeset
   175
"aval' (N i) S = num' i" |
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   176
"aval' (V x) S = S x" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   177
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   178
51807
nipkow
parents: 51791
diff changeset
   179
definition "asem x e S = (case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S)))"
51694
nipkow
parents: 51628
diff changeset
   180
51807
nipkow
parents: 51791
diff changeset
   181
definition "step' = Step asem (\<lambda>b S. S)"
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51359
diff changeset
   182
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   183
lemma strip_step'[simp]: "strip(step' S C) = strip C"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   184
by(simp add: step'_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   185
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   186
definition AI :: "com \<Rightarrow> 'av st option acom option" where
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   187
"AI c = pfp (step' \<top>) (bot c)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   188
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   189
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   190
abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   191
where "\<gamma>\<^sub>s == \<gamma>_fun \<gamma>"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   192
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   193
abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   194
where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   195
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   196
abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   197
where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   198
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   199
lemma gamma_s_Top[simp]: "\<gamma>\<^sub>s \<top> = UNIV"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   200
by(simp add: top_fun_def \<gamma>_fun_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   201
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   202
lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   203
by (simp add: top_option_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   204
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   205
lemma mono_gamma_s: "f1 \<le> f2 \<Longrightarrow> \<gamma>\<^sub>s f1 \<subseteq> \<gamma>\<^sub>s f2"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   206
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   207
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   208
lemma mono_gamma_o:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   209
  "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   210
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   211
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   212
lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   213
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   214
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   215
text\<open>Correctness:\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   216
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   217
lemma aval'_correct: "s \<in> \<gamma>\<^sub>s S \<Longrightarrow> aval a s \<in> \<gamma>(aval' a S)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   218
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   219
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   220
lemma in_gamma_update: "\<lbrakk> s \<in> \<gamma>\<^sub>s S; i \<in> \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) \<in> \<gamma>\<^sub>s(S(x := a))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   221
by(simp add: \<gamma>_fun_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   222
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   223
lemma gamma_Step_subcomm:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   224
  assumes "\<And>x e S. f1 x e (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (f2 x e S)"  "\<And>b S. g1 b (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (g2 b S)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   225
  shows "Step f1 g1 (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (Step f2 g2 S C)"
54944
nipkow
parents: 53015
diff changeset
   226
by (induction C arbitrary: S) (auto simp: mono_gamma_o assms)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   227
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   228
lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   229
unfolding step_def step'_def
51694
nipkow
parents: 51628
diff changeset
   230
by(rule gamma_Step_subcomm)
51974
9c80e62161a5 tuned names
nipkow
parents: 51826
diff changeset
   231
  (auto simp: aval'_correct in_gamma_update asem_def split: option.splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   232
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   233
lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   234
proof(simp add: CS_def AI_def)
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49396
diff changeset
   235
  assume 1: "pfp (step' \<top>) (bot c) = Some C"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   236
  have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
   237
  have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C"  \<comment> \<open>transfer the pfp'\<close>
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
   238
  proof(rule order_trans)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   239
    show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)" by(rule step_step')
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   240
    show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp'])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   241
  qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   242
  have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   243
  have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   244
    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 2])
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   245
  thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   246
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   247
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   248
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   249
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   250
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   251
subsubsection "Monotonicity"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   252
51721
nipkow
parents: 51710
diff changeset
   253
locale Abs_Int_fun_mono = Abs_Int_fun +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   254
assumes mono_plus': "a1 \<le> b1 \<Longrightarrow> a2 \<le> b2 \<Longrightarrow> plus' a1 a2 \<le> plus' b1 b2"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   255
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   256
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   257
lemma mono_aval': "S \<le> S' \<Longrightarrow> aval' e S \<le> aval' e S'"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   258
by(induction e)(auto simp: le_fun_def mono_plus')
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   259
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   260
lemma mono_update: "a \<le> a' \<Longrightarrow> S \<le> S' \<Longrightarrow> S(x := a) \<le> S'(x := a')"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   261
by(simp add: le_fun_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   262
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   263
lemma mono_step': "S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> step' S1 C1 \<le> step' S2 C2"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   264
unfolding step'_def
51694
nipkow
parents: 51628
diff changeset
   265
by(rule mono2_Step)
51807
nipkow
parents: 51791
diff changeset
   266
  (auto simp: mono_update mono_aval' asem_def split: option.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   267
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   268
lemma mono_step'_top: "C \<le> C' \<Longrightarrow> step' \<top> C \<le> step' \<top> C'"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   269
by (metis mono_step' order_refl)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   270
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   271
lemma AI_least_pfp: assumes "AI c = Some C" "step' \<top> C' \<le> C'" "strip C' = c"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   272
shows "C \<le> C'"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   273
by(rule pfp_bot_least[OF _ _ assms(2,3) assms(1)[unfolded AI_def]])
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   274
  (simp_all add: mono_step'_top)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   275
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   276
end
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   277
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   278
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   279
instantiation acom :: (type) vars
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   280
begin
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   281
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   282
definition "vars_acom = vars o strip"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   283
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   284
instance ..
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   285
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   286
end
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   287
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   288
lemma finite_Cvars: "finite(vars(C::'a acom))"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   289
by(simp add: vars_acom_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   290
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   291
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   292
subsubsection "Termination"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   293
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   294
lemma pfp_termination:
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   295
fixes x0 :: "'a::order" and m :: "'a \<Rightarrow> nat"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   296
assumes mono: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   297
and m: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x < y \<Longrightarrow> m x > m y"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   298
and I: "\<And>x y. I x \<Longrightarrow> I(f x)" and "I x0" and "x0 \<le> f x0"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   299
shows "\<exists>x. pfp f x0 = Some x"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   300
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \<le> f x"])
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   301
  show "wf {(y,x). ((I x \<and> x \<le> f x) \<and> \<not> f x \<le> x) \<and> y = f x}"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   302
    by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   303
next
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   304
  show "I x0 \<and> x0 \<le> f x0" using \<open>I x0\<close> \<open>x0 \<le> f x0\<close> by blast
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   305
next
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   306
  fix x assume "I x \<and> x \<le> f x" thus "I(f x) \<and> f x \<le> f(f x)"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   307
    by (blast intro: I mono)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   308
qed
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   309
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   310
lemma le_iff_le_annos: "C1 \<le> C2 \<longleftrightarrow>
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   311
  strip C1 = strip C2 \<and> (\<forall> i<size(annos C1). annos C1 ! i \<le> annos C2 ! i)"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   312
by(simp add: less_eq_acom_def anno_def)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   313
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   314
locale Measure1_fun =
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   315
fixes m :: "'av::top \<Rightarrow> nat"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   316
fixes h :: "nat"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   317
assumes h: "m x \<le> h"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   318
begin
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   319
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   320
definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   321
"m_s S X = (\<Sum> x \<in> X. m(S x))"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   322
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   323
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   324
by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h])
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   325
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   326
fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   327
"m_o (Some S) X = m_s S X" |
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   328
"m_o None X = h * card X + 1"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   329
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   330
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (h*card X + 1)"
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   331
by(cases opt)(auto simp add: m_s_h le_SucI dest: m_s_h)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   332
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   333
definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^sub>c") where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61179
diff changeset
   334
"m_c C = sum_list (map (\<lambda>a. m_o a (vars C)) (annos C))"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   335
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   336
text\<open>Upper complexity bound:\<close>
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   337
lemma m_c_h: "m_c C \<le> size(annos C) * (h * card(vars C) + 1)"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   338
proof-
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   339
  let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)"
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   340
  have "m_c C = (\<Sum>i<?a. m_o (annos C ! i) ?X)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   341
    by(simp add: m_c_def sum_list_sum_nth atLeast0LessThan)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   342
  also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   343
    apply(rule sum_mono) using m_o_h[OF finite_Cvars] by simp
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   344
  also have "\<dots> = ?a * (h * ?n + 1)" by simp
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   345
  finally show ?thesis .
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   346
qed
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   347
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   348
end
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   349
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   350
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   351
locale Measure_fun = Measure1_fun where m=m
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   352
  for m :: "'av::semilattice_sup_top \<Rightarrow> nat" +
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   353
assumes m2: "x < y \<Longrightarrow> m x > m y"
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   354
begin
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   355
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68778
diff changeset
   356
text\<open>The predicates \<open>top_on_ty a X\<close> that follow describe that any abstract
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68778
diff changeset
   357
state in \<open>a\<close> maps all variables in \<open>X\<close> to @{term \<top>}.
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   358
This is an important invariant for the termination proof where we argue that only
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   359
the finitely many variables in the program change. That the others do not change
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   360
follows because they remain @{term \<top>}.\<close>
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   361
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   362
fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   363
"top_on_st S X = (\<forall>x\<in>X. S x = \<top>)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   364
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   365
fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   366
"top_on_opt (Some S) X = top_on_st S X" |
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   367
"top_on_opt None X = True"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   368
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   369
definition top_on_acom :: "'av st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   370
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   371
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   372
lemma top_on_top: "top_on_opt \<top> X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   373
by(auto simp: top_option_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   374
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   375
lemma top_on_bot: "top_on_acom (bot c) X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   376
by(auto simp add: top_on_acom_def bot_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   377
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   378
lemma top_on_post: "top_on_acom C X \<Longrightarrow> top_on_opt (post C) X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   379
by(simp add: top_on_acom_def post_in_annos)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   380
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   381
lemma top_on_acom_simps:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   382
  "top_on_acom (SKIP {Q}) X = top_on_opt Q X"
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   383
  "top_on_acom (x ::= e {Q}) X = top_on_opt Q X"
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 52022
diff changeset
   384
  "top_on_acom (C1;;C2) X = (top_on_acom C1 X \<and> top_on_acom C2 X)"
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   385
  "top_on_acom (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) X =
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   386
   (top_on_opt P1 X \<and> top_on_acom C1 X \<and> top_on_opt P2 X \<and> top_on_acom C2 X \<and> top_on_opt Q X)"
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   387
  "top_on_acom ({I} WHILE b DO {P} C {Q}) X =
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   388
   (top_on_opt I X \<and> top_on_acom C X \<and> top_on_opt P X \<and> top_on_opt Q X)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   389
by(auto simp add: top_on_acom_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   390
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   391
lemma top_on_sup:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   392
  "top_on_opt o1 X \<Longrightarrow> top_on_opt o2 X \<Longrightarrow> top_on_opt (o1 \<squnion> o2) X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   393
apply(induction o1 o2 rule: sup_option.induct)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   394
apply(auto)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   395
done
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   396
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   397
lemma top_on_Step: fixes C :: "'av st option acom"
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   398
assumes "!!x e S. \<lbrakk>top_on_opt S X; x \<notin> X; vars e \<subseteq> -X\<rbrakk> \<Longrightarrow> top_on_opt (f x e S) X"
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   399
        "!!b S. top_on_opt S X \<Longrightarrow> vars b \<subseteq> -X \<Longrightarrow> top_on_opt (g b S) X"
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   400
shows "\<lbrakk> vars C \<subseteq> -X; top_on_opt S X; top_on_acom C X \<rbrakk> \<Longrightarrow> top_on_acom (Step f g S C) X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   401
proof(induction C arbitrary: S)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   402
qed (auto simp: top_on_acom_simps vars_acom_def top_on_post top_on_sup assms)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   403
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   404
lemma m1: "x \<le> y \<Longrightarrow> m x \<ge> m y"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   405
by(auto simp: le_less m2)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   406
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   407
lemma m_s2_rep: assumes "finite(X)" and "S1 = S2 on -X" and "\<forall>x. S1 x \<le> S2 x" and "S1 \<noteq> S2"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   408
shows "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   409
proof-
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   410
  from assms(3) have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S2 x)" by (simp add: m1)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   411
  from assms(2,3,4) have "\<exists>x\<in>X. S1 x < S2 x"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   412
    by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   413
  hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   414
  from sum_strict_mono_ex1[OF \<open>finite X\<close> 1 2]
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   415
  show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" .
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   416
qed
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   417
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   418
lemma m_s2: "finite(X) \<Longrightarrow> S1 = S2 on -X \<Longrightarrow> S1 < S2 \<Longrightarrow> m_s S1 X > m_s S2 X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   419
apply(auto simp add: less_fun_def m_s_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   420
apply(simp add: m_s2_rep le_fun_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   421
done
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   422
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   423
lemma m_o2: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow>
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   424
  o1 < o2 \<Longrightarrow> m_o o1 X > m_o o2 X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   425
proof(induction o1 o2 rule: less_eq_option.induct)
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   426
  case 1 thus ?case by (auto simp: m_s2 less_option_def)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   427
next
51749
c27bb7994bd3 moved defs into locale to reduce unnecessary polymorphism; tuned
nipkow
parents: 51722
diff changeset
   428
  case 2 thus ?case by(auto simp: less_option_def le_imp_less_Suc m_s_h)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   429
next
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   430
  case 3 thus ?case by (auto simp: less_option_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   431
qed
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   432
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   433
lemma m_o1: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow>
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   434
  o1 \<le> o2 \<Longrightarrow> m_o o1 X \<ge> m_o o2 X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   435
by(auto simp: le_less m_o2)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   436
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   437
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   438
lemma m_c2: "top_on_acom C1 (-vars C1) \<Longrightarrow> top_on_acom C2 (-vars C2) \<Longrightarrow>
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   439
  C1 < C2 \<Longrightarrow> m_c C1 > m_c C2"
51783
f4a00cdae743 simplified def
nipkow
parents: 51754
diff changeset
   440
proof(auto simp add: le_iff_le_annos size_annos_same[of C1 C2] vars_acom_def less_acom_def)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   441
  let ?X = "vars(strip C2)"
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   442
  assume top: "top_on_acom C1 (- vars(strip C2))"  "top_on_acom C2 (- vars(strip C2))"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   443
  and strip_eq: "strip C1 = strip C2"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   444
  and 0: "\<forall>i<size(annos C2). annos C1 ! i \<le> annos C2 ! i"
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   445
  hence 1: "\<forall>i<size(annos C2). m_o (annos C1 ! i) ?X \<ge> m_o (annos C2 ! i) ?X"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   446
    apply (auto simp: all_set_conv_all_nth vars_acom_def top_on_acom_def)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   447
    by (metis (lifting, no_types) finite_cvars m_o1 size_annos_same2)
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   448
  fix i assume i: "i < size(annos C2)" "\<not> annos C2 ! i \<le> annos C1 ! i"
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   449
  have topo1: "top_on_opt (annos C1 ! i) (- ?X)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   450
    using i(1) top(1) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   451
  have topo2: "top_on_opt (annos C2 ! i) (- ?X)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   452
    using i(1) top(2) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   453
  from i have "m_o (annos C1 ! i) ?X > m_o (annos C2 ! i) ?X" (is "?P i")
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   454
    by (metis 0 less_option_def m_o2[OF finite_cvars topo1] topo2)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   455
  hence 2: "\<exists>i < size(annos C2). ?P i" using \<open>i < size(annos C2)\<close> by blast
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   456
  have "(\<Sum>i<size(annos C2). m_o (annos C2 ! i) ?X)
c4db685eaed0 more standard argument order
nipkow
parents: 51785
diff changeset
   457
         < (\<Sum>i<size(annos C2). m_o (annos C1 ! i) ?X)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   458
    apply(rule sum_strict_mono_ex1) using 1 2 by (auto)
51783
f4a00cdae743 simplified def
nipkow
parents: 51754
diff changeset
   459
  thus ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   460
    by(simp add: m_c_def vars_acom_def strip_eq sum_list_sum_nth atLeast0LessThan size_annos_same[OF strip_eq])
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   461
qed
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   462
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   463
end
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   464
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   465
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   466
locale Abs_Int_fun_measure =
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   467
  Abs_Int_fun_mono where \<gamma>=\<gamma> + Measure_fun where m=m
51826
054a40461449 canonical names of classes
nipkow
parents: 51807
diff changeset
   468
  for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   469
begin
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   470
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   471
lemma top_on_step': "top_on_acom C (-vars C) \<Longrightarrow> top_on_acom (step' \<top> C) (-vars C)"
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   472
unfolding step'_def
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   473
by(rule top_on_Step)
51807
nipkow
parents: 51791
diff changeset
   474
  (auto simp add: top_option_def asem_def split: option.splits)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   475
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   476
lemma AI_Some_measure: "\<exists>C. AI c = Some C"
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   477
unfolding AI_def
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51783
diff changeset
   478
apply(rule pfp_termination[where I = "\<lambda>C. top_on_acom C (- vars C)" and m="m_c"])
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   479
apply(simp_all add: m_c2 mono_step'_top bot_least top_on_bot)
51754
nipkow
parents: 51749
diff changeset
   480
using top_on_step' apply(auto simp add: vars_acom_def)
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   481
done
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51721
diff changeset
   482
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   483
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   484
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   485
text\<open>Problem: not executable because of the comparison of abstract states,
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 64267
diff changeset
   486
i.e. functions, in the pre-fixpoint computation.\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   487
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   488
end