src/HOL/IMP/Abs_Int1.thy
author nipkow
Wed, 17 Apr 2013 21:23:35 +0200
changeset 51712 30624dab6054
parent 51711 df3426139651
child 51722 3da99469cc1b
permissions -rw-r--r--
tuned
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
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     3
theory Abs_Int1
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     4
imports Abs_State
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     5
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     6
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
     7
instantiation acom :: (type) vars
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
     8
begin
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
     9
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    10
definition "vars_acom = vars o strip"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    11
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    12
instance ..
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    13
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    14
end
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    15
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    16
lemma finite_Cvars: "finite(vars(C::'a acom))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    17
by(simp add: vars_acom_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    18
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    19
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    20
lemma le_iff_le_annos_zip: "C1 \<le> C2 \<longleftrightarrow>
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    21
 (\<forall> (a1,a2) \<in> set(zip (annos C1) (annos C2)). a1 \<le> a2) \<and> strip C1 = strip C2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    22
by(induct C1 C2 rule: less_eq_acom.induct) (auto simp: size_annos_same2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    23
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    24
lemma le_iff_le_annos: "C1 \<le> C2 \<longleftrightarrow>
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    25
  strip C1 = strip C2 \<and> (\<forall> i<size(annos C1). annos C1 ! i \<le> annos C2 ! i)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    26
by(auto simp add: le_iff_le_annos_zip set_zip) (metis size_annos_same2)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    27
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    28
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    29
subsection "Computable Abstract Interpretation"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    30
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    31
text{* Abstract interpretation over type @{text st} instead of functions. *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    32
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    33
context Gamma
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    34
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    35
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    36
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
50896
nipkow
parents: 49547
diff changeset
    37
"aval' (N i) S = num' i" |
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    38
"aval' (V x) S = fun S x" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    39
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    40
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    41
lemma aval'_sound: "s : \<gamma>\<^isub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    42
by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    43
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    44
lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    45
  assumes "!!x e S. f1 x e (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (f2 x e S)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    46
          "!!b S. g1 b (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (g2 b S)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    47
  shows "Step f1 g1 (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (Step f2 g2 S C)"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    48
proof(induction C arbitrary: S)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    49
qed (auto simp: assms intro!: mono_gamma_o sup_ge1 sup_ge2)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    50
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    51
lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(update S x a)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    52
by(simp add: \<gamma>_st_def)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    53
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    54
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    55
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    56
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    57
text{* The for-clause (here and elsewhere) only serves the purpose of fixing
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    58
the name of the type parameter @{typ 'av} which would otherwise be renamed to
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    59
@{typ 'a}. *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    60
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    61
locale Abs_Int = Gamma where \<gamma>=\<gamma> for \<gamma> :: "'av::semilattice \<Rightarrow> val set"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    62
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    63
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    64
definition "step' = Step
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    65
  (\<lambda>x e S. case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S)))
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    66
  (\<lambda>b S. S)"
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    67
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    68
definition AI :: "com \<Rightarrow> 'av st option acom option" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    69
"AI c = pfp (step' \<top>) (bot c)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    70
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    71
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    72
lemma strip_step'[simp]: "strip(step' S C) = strip C"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    73
by(simp add: step'_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    74
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    75
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    76
text{* Soundness: *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    77
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    78
lemma step_step': "step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    79
unfolding step_def step'_def
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    80
by(rule gamma_Step_subcomm)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    81
  (auto simp: intro!: aval'_sound in_gamma_update split: option.splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    82
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
    83
lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    84
proof(simp add: CS_def AI_def)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    85
  assume 1: "pfp (step' \<top>) (bot c) = Some C"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    86
  have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    87
  have 2: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"  --"transfer the pfp'"
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
    88
  proof(rule order_trans)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    89
    show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" by(rule step_step')
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    90
    show "... \<le> \<gamma>\<^isub>c C" by (metis mono_gamma_c[OF pfp'])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    91
  qed
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    92
  have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    93
  have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    94
    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 2])
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
    95
  thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    96
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    97
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    98
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    99
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   100
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   101
subsubsection "Monotonicity"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   102
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   103
lemma le_sup_disj: "x \<le> y \<or> x \<le> z \<Longrightarrow> x \<le> y \<squnion> (z::_::semilattice_sup)"
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
   104
by (metis sup_ge1 sup_ge2 order_trans)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   105
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   106
theorem mono2_Step: fixes C1 :: "'a::semilattice acom"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   107
  assumes "!!x e S1 S2. S1 \<le> S2 \<Longrightarrow> f x e S1 \<le> f x e S2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   108
          "!!b S1 S2. S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   109
  shows "S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   110
by(induction C1 C2 arbitrary: S1 S2 rule: less_eq_acom.induct)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   111
  (auto simp: mono_post assms le_sup_disj)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   112
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   113
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   114
locale Abs_Int_mono = Abs_Int +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   115
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
   116
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   117
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   118
lemma mono_aval': "S1 \<le> S2 \<Longrightarrow> aval' e S1 \<le> aval' e S2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   119
by(induction e) (auto simp: mono_plus' mono_fun)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   120
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   121
theorem 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
   122
unfolding step'_def
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   123
by(rule mono2_Step) (auto simp: mono_aval' split: option.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   124
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   125
lemma mono_step'_top: "C \<le> C' \<Longrightarrow> step' \<top> C \<le> step' \<top> C'"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   126
by (metis mono_step' order_refl)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   127
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   128
lemma AI_least_pfp: assumes "AI c = Some C" "step' \<top> C' \<le> C'" "strip C' = c"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   129
shows "C \<le> C'"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   130
by(rule pfp_bot_least[OF _ _ assms(2,3) assms(1)[unfolded AI_def]])
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   131
  (simp_all add: mono_step'_top)
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   132
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   133
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   134
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   135
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   136
subsubsection "Termination"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   137
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   138
lemma pfp_termination:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   139
fixes x0 :: "'a::order" and m :: "'a \<Rightarrow> nat"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   140
assumes mono: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   141
and m: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x < y \<Longrightarrow> m x > m y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   142
and I: "\<And>x y. I x \<Longrightarrow> I(f x)" and "I x0" and "x0 \<le> f x0"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   143
shows "\<exists>x. pfp f x0 = Some x"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   144
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \<le> f x"])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   145
  show "wf {(y,x). ((I x \<and> x \<le> f x) \<and> \<not> f x \<le> x) \<and> y = f x}"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   146
    by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   147
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   148
  show "I x0 \<and> x0 \<le> f x0" using `I x0` `x0 \<le> f x0` by blast
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   149
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   150
  fix x assume "I x \<and> x \<le> f x" thus "I(f x) \<and> f x \<le> f(f x)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   151
    by (blast intro: I mono)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   152
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   153
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   154
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   155
locale Measure1 =
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   156
fixes m :: "'av::{order,top} \<Rightarrow> nat"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   157
fixes h :: "nat"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   158
assumes h: "m x \<le> h"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   159
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   160
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   161
definition m_s :: "vname set \<Rightarrow> 'av st \<Rightarrow> nat" ("m\<^isub>s") where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   162
"m_s X S = (\<Sum> x \<in> X. m(fun S x))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   163
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   164
lemma m_s_h: "finite X \<Longrightarrow> m_s X S \<le> h * card X"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   165
by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   166
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   167
definition m_o :: "vname set \<Rightarrow> 'av st option \<Rightarrow> nat" ("m\<^isub>o") where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   168
"m_o X opt = (case opt of None \<Rightarrow> h * card X + 1 | Some S \<Rightarrow> m_s X S)"
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   169
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   170
lemma m_o_h: "finite X \<Longrightarrow> m_o X opt \<le> (h*card X + 1)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   171
by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   172
49432
3f4104ccca77 beautified names
nipkow
parents: 49431
diff changeset
   173
definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^isub>c") where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   174
"m_c C = (\<Sum>i<size(annos C). m_o (vars C) (annos C ! i))"
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   175
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   176
text{* Upper complexity bound: *}
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   177
lemma m_c_h: "m_c C \<le> size(annos C) * (h * card(vars C) + 1)"
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   178
proof-
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   179
  let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   180
  have "m_c C = (\<Sum>i<?a. m_o ?X (annos C ! i))" by(simp add: m_c_def)
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   181
  also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   182
    apply(rule setsum_mono) using m_o_h[OF finite_Cvars] by simp
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   183
  also have "\<dots> = ?a * (h * ?n + 1)" by simp
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   184
  finally show ?thesis .
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   185
qed
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   186
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   187
end
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   188
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   189
text{* The predicates @{text "top_on X a"} that follow describe that @{text a} is some object
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   190
that maps all variables in @{text X} to @{term \<top>}.
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   191
This is an important invariant for the termination proof where we argue that only
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   192
the finitely many variables in the program change. That the others do not change
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   193
follows because they remain @{term \<top>}. *}
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   194
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   195
class top_on =
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   196
fixes top_on :: "vname set \<Rightarrow> 'a \<Rightarrow> bool"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   197
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   198
instantiation st :: (top)top_on
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   199
begin
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   200
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   201
fun top_on_st :: "vname set \<Rightarrow> 'a st \<Rightarrow> bool" where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   202
"top_on_st X F = (\<forall>x\<in>X. fun F x = \<top>)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   203
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   204
instance ..
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   205
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   206
end
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   207
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   208
instantiation option :: (top_on)top_on
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   209
begin
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   210
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   211
fun top_on_option :: "vname set \<Rightarrow> 'a option \<Rightarrow> bool" where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   212
"top_on_option X (Some F) = top_on X F" |
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   213
"top_on_option X None = True"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   214
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   215
instance ..
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   216
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   217
end
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   218
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   219
instantiation acom :: (top_on)top_on
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   220
begin
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   221
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   222
definition top_on_acom :: "vname set \<Rightarrow> 'a acom \<Rightarrow> bool" where
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   223
"top_on_acom X C = (\<forall>a \<in> set(annos C). top_on X a)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   224
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   225
instance ..
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   226
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   227
end
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   228
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   229
lemma top_on_top: "top_on X (\<top>::_ st option)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   230
by(auto simp: top_option_def fun_top)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   231
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   232
lemma top_on_bot: "top_on X (bot c)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   233
by(auto simp add: top_on_acom_def bot_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   234
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   235
lemma top_on_post: "top_on X C \<Longrightarrow> top_on X (post C)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   236
by(simp add: top_on_acom_def post_in_annos)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   237
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   238
lemma top_on_acom_simps:
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   239
  "top_on X (SKIP {Q}) = top_on X Q"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   240
  "top_on X (x ::= e {Q}) = top_on X Q"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   241
  "top_on X (C1;C2) = (top_on X C1 \<and> top_on X C2)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   242
  "top_on X (IF b THEN {P1} C1 ELSE {P2} C2 {Q}) =
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   243
   (top_on X P1 \<and> top_on X C1 \<and> top_on X P2 \<and> top_on X C2 \<and> top_on X Q)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   244
  "top_on X ({I} WHILE b DO {P} C {Q}) =
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   245
   (top_on X I \<and> top_on X C \<and> top_on X P \<and> top_on X Q)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   246
by(auto simp add: top_on_acom_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   247
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   248
lemma top_on_sup:
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   249
  "top_on X o1 \<Longrightarrow> top_on X o2 \<Longrightarrow> top_on X (o1 \<squnion> o2 :: _ st option)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   250
apply(induction o1 o2 rule: sup_option.induct)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   251
apply(auto)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   252
by transfer simp
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   253
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   254
lemma top_on_Step: fixes C :: "('a::semilattice)st option acom"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   255
assumes "!!x e S. \<lbrakk>top_on X S; x \<notin> X; vars e \<subseteq> -X\<rbrakk> \<Longrightarrow> top_on X (f x e S)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   256
        "!!b S. top_on X S \<Longrightarrow> vars b \<subseteq> -X \<Longrightarrow> top_on X (g b S)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   257
shows "\<lbrakk> vars C \<subseteq> -X; top_on X S; top_on X C \<rbrakk> \<Longrightarrow> top_on X (Step f g S C)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   258
proof(induction C arbitrary: S)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   259
qed (auto simp: top_on_acom_simps vars_acom_def top_on_post top_on_sup assms)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   260
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   261
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   262
locale Measure = Measure1 +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   263
assumes m2: "x < y \<Longrightarrow> m x > m y"
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   264
begin
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   265
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   266
lemma m1: "x \<le> y \<Longrightarrow> m x \<ge> m y"
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   267
by(auto simp: le_less m2)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   268
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   269
lemma m_s2_rep: assumes "finite(X)" and "S1 = S2 on -X" and "\<forall>x. S1 x \<le> S2 x" and "S1 \<noteq> S2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   270
shows "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   271
proof-
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   272
  from assms(3) have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S2 x)" by (simp add: m1)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   273
  from assms(2,3,4) have "EX x:X. S1 x < S2 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   274
    by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   275
  hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   276
  from setsum_strict_mono_ex1[OF `finite X` 1 2]
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   277
  show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" .
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   278
qed
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   279
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   280
lemma m_s2: "finite(X) \<Longrightarrow> fun S1 = fun S2 on -X \<Longrightarrow> S1 < S2 \<Longrightarrow> m_s X S1 > m_s X S2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   281
apply(auto simp add: less_st_def m_s_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   282
apply (transfer fixing: m)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   283
apply(simp add: less_eq_st_rep_iff eq_st_def m_s2_rep)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   284
done
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   285
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   286
lemma m_o2: "finite X \<Longrightarrow> top_on (-X) o1 \<Longrightarrow> top_on (-X) o2 \<Longrightarrow>
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   287
  o1 < o2 \<Longrightarrow> m_o X o1 > m_o X o2"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   288
proof(induction o1 o2 rule: less_eq_option.induct)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   289
  case 1 thus ?case by (auto simp: m_o_def m_s2 less_option_def)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   290
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   291
  case 2 thus ?case by(auto simp: m_o_def less_option_def le_imp_less_Suc m_s_h)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   292
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   293
  case 3 thus ?case by (auto simp: less_option_def)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   294
qed
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   295
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   296
lemma m_o1: "finite X \<Longrightarrow> top_on (-X) o1 \<Longrightarrow> top_on (-X) o2 \<Longrightarrow>
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   297
  o1 \<le> o2 \<Longrightarrow> m_o X o1 \<ge> m_o X o2"
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   298
by(auto simp: le_less m_o2)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   299
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   300
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   301
lemma m_c2: "top_on (-vars C1) C1 \<Longrightarrow> top_on (-vars C2) C2 \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   302
  C1 < C2 \<Longrightarrow> m_c C1 > m_c C2"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   303
proof(auto simp add: le_iff_le_annos m_c_def size_annos_same[of C1 C2] vars_acom_def less_acom_def)
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   304
  let ?X = "vars(strip C2)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   305
  assume top: "top_on (- vars(strip C2)) C1"  "top_on (- vars(strip C2)) C2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   306
  and strip_eq: "strip C1 = strip C2"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   307
  and 0: "\<forall>i<size(annos C2). annos C1 ! i \<le> annos C2 ! i"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   308
  hence 1: "\<forall>i<size(annos C2). m_o ?X (annos C1 ! i) \<ge> m_o ?X (annos C2 ! i)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   309
    apply (auto simp: all_set_conv_all_nth vars_acom_def top_on_acom_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   310
    by (metis finite_cvars m_o1 size_annos_same2)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   311
  fix i assume i: "i < size(annos C2)" "\<not> annos C2 ! i \<le> annos C1 ! i"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   312
  have topo1: "top_on (- ?X) (annos C1 ! i)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   313
    using i(1) top(1) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   314
  have topo2: "top_on (- ?X) (annos C2 ! i)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   315
    using i(1) top(2) by(simp add: top_on_acom_def size_annos_same[OF strip_eq])
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   316
  from i have "m_o ?X (annos C1 ! i) > m_o ?X (annos C2 ! i)" (is "?P i")
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   317
    by (metis 0 less_option_def m_o2[OF finite_cvars topo1] topo2)
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   318
  hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   319
  show "(\<Sum>i<size(annos C2). m_o ?X (annos C2 ! i))
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   320
         < (\<Sum>i<size(annos C2). m_o ?X (annos C1 ! i))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   321
    apply(rule setsum_strict_mono_ex1) using 1 2 by (auto)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   322
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   323
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   324
end
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   325
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   326
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   327
locale Abs_Int_measure =
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   328
  Abs_Int_mono where \<gamma>=\<gamma> + Measure where m=m
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   329
  for \<gamma> :: "'av::semilattice \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   330
begin
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   331
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   332
lemma top_on_step': "\<lbrakk> vars C \<subseteq> -X; top_on X C \<rbrakk> \<Longrightarrow> top_on X (step' \<top> C)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   333
unfolding step'_def
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   334
by(rule top_on_Step)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   335
  (auto simp add: top_option_def fun_top split: option.splits)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   336
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   337
lemma AI_Some_measure: "\<exists>C. AI c = Some C"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   338
unfolding AI_def
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   339
apply(rule pfp_termination[where I = "\<lambda>C. strip C = c \<and> top_on (- vars C) C" and m="m_c"])
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   340
apply(simp_all add: m_c2 mono_step'_top bot_least top_on_bot)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   341
apply(auto simp add: top_on_step' vars_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   342
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   343
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   344
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   345
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   346
end