src/HOL/IMP/AbsInt2.thy
author nipkow
Wed, 21 Sep 2011 07:04:04 +0200
changeset 45022 3c888c58e10b
parent 45020 21334181f820
child 45023 76abd26e2e2d
permissions -rw-r--r--
Added proofs about narowing
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     2
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     3
theory AbsInt2
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     4
imports AbsInt1_ivl
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     5
begin
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
     6
45022
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
     7
context preord
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
     8
begin
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
     9
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    10
definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    11
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    12
lemma monoD: "mono f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" by(simp add: mono_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    13
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    14
lemma mono_comp: "mono f \<Longrightarrow> mono g \<Longrightarrow> mono (g o f)"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    15
by(simp add: mono_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    16
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    17
declare le_trans[trans]
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    18
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    19
end
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    20
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    21
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    22
subsection "Widening and Narrowing"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    23
45022
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    24
text{* Jumping to the trivial post-fixed point @{const Top} in case @{text k}
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    25
rounds of iteration did not reach a post-fixed point (as in @{const iter}) is
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    26
a trivial widening step. We generalise this idea and complement it with
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    27
narrowing, a process to regain precision.
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    28
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    29
Class @{text WN} makes some assumptions about the widening and narrowing
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    30
operators. The assumptions serve two purposes. Together with a further
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    31
assumption that certain chains become stationary, they permit to prove
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    32
termination of the fixed point iteration, which we do not --- we limit the
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    33
number of iterations as before. The second purpose of the narrowing
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    34
assumptions is to prove that the narrowing iteration keeps on producing
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    35
post-fixed points and that it goes down. However, this requires the function
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    36
being iterated to be monotone. Unfortunately, abstract interpretation with
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    37
widening is not monotone. Hence the (recursive) abstract interpretation of a
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    38
loop body that again contains a loop may result in a non-monotone
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    39
function. Therefore our narrowing interation needs to check at every step
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    40
that a post-fixed point is maintained, and we cannot prove that the precision
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    41
increases. *}
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    42
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    43
class WN = SL_top +
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    44
fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    45
assumes widen: "y \<sqsubseteq> x \<nabla> y"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    46
fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    47
assumes narrow1: "y \<sqsubseteq> x \<Longrightarrow> y \<sqsubseteq> x \<triangle> y"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    48
assumes narrow2: "y \<sqsubseteq> x \<Longrightarrow> x \<triangle> y \<sqsubseteq> x"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    49
begin
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    50
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    51
fun iter_up :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    52
"iter_up f 0 x = Top" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    53
"iter_up f (Suc n) x =
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    54
  (let fx = f x in if fx \<sqsubseteq> x then x else iter_up f n (x \<nabla> fx))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    55
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    56
lemma iter_up_pfp: "f(iter_up f n x) \<sqsubseteq> iter_up f n x"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
    57
apply (induction n arbitrary: x)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    58
 apply (simp)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    59
apply (simp add: Let_def)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    60
done
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    61
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    62
fun iter_down :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    63
"iter_down f 0 x = x" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    64
"iter_down f (Suc n) x =
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    65
  (let y = x \<triangle> f x in if f y \<sqsubseteq> y then iter_down f n y else x)"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    66
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    67
lemma iter_down_pfp: "f x \<sqsubseteq> x \<Longrightarrow> f(iter_down f n x) \<sqsubseteq> iter_down f n x"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44944
diff changeset
    68
apply (induction n arbitrary: x)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    69
 apply (simp)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    70
apply (simp add: Let_def)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    71
done
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    72
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    73
definition iter' :: "nat \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    74
"iter' m n f x =
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    75
  (let f' = (\<lambda>y. x \<squnion> f y) in iter_down f' n (iter_up f' m x))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    76
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    77
lemma iter'_pfp_above:
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    78
shows "f(iter' m n f x0) \<sqsubseteq> iter' m n f x0"
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    79
and "x0 \<sqsubseteq> iter' m n f x0"
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    80
using iter_up_pfp[of "\<lambda>x. x0 \<squnion> f x"] iter_down_pfp[of "\<lambda>x. x0 \<squnion> f x"]
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
    81
by(auto simp add: iter'_def Let_def)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
    82
45022
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    83
text{* This is how narrowing works on monotone functions: you just iterate. *}
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    84
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    85
abbreviation iter_down_mono :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    86
"iter_down_mono f n x == ((\<lambda>x. x \<triangle> f x)^^n) x"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    87
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    88
text{* Narrowing always yields a post-fixed point: *}
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    89
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    90
lemma iter_down_mono_pfp: assumes "mono f" and "f x0 \<sqsubseteq> x0" 
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    91
defines "x n == iter_down_mono f n x0"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    92
shows "f(x n) \<sqsubseteq> x n"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    93
proof (induction n)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    94
  case 0 show ?case by (simp add: x_def assms(2))
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    95
next
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    96
  case (Suc n)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    97
  have "f (x (Suc n)) = f(x n \<triangle> f(x n))" by(simp add: x_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    98
  also have "\<dots> \<sqsubseteq> f(x n)" by(rule monoD[OF `mono f` narrow2[OF Suc]])
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
    99
  also have "\<dots> \<sqsubseteq> x n \<triangle> f(x n)" by(rule narrow1[OF Suc])
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   100
  also have "\<dots> = x(Suc n)" by(simp add: x_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   101
  finally show ?case .
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   102
qed
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   103
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   104
text{* Narrowing can only increase precision: *}
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   105
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   106
lemma iter_down_down: assumes "mono f" and "f x0 \<sqsubseteq> x0" 
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   107
defines "x n == iter_down_mono f n x0"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   108
shows "x n \<sqsubseteq> x0"
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   109
proof (induction n)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   110
  case 0 show ?case by(simp add: x_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   111
next
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   112
  case (Suc n)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   113
  have "x(Suc n) = x n \<triangle> f(x n)" by(simp add: x_def)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   114
  also have "\<dots> \<sqsubseteq> x n" unfolding x_def
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   115
    by(rule narrow2[OF iter_down_mono_pfp[OF assms(1), OF assms(2)]])
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   116
  also have "\<dots> \<sqsubseteq> x0" by(rule Suc)
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   117
  finally show ?case .
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   118
qed
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   119
3c888c58e10b Added proofs about narowing
nipkow
parents: 45020
diff changeset
   120
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   121
end
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   122
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   123
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   124
instantiation ivl :: WN
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   125
begin
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   126
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   127
definition "widen_ivl ivl1 ivl2 =
45019
nipkow
parents: 45018
diff changeset
   128
  ((*if is_empty ivl1 then ivl2 else if is_empty ivl2 then ivl1 else*)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   129
     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   130
       I (if le_option False l2 l1 \<and> l2 \<noteq> l1 then None else l2)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   131
         (if le_option True h1 h2 \<and> h1 \<noteq> h2 then None else h2))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   132
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   133
definition "narrow_ivl ivl1 ivl2 =
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   134
  ((*if is_empty ivl1 \<or> is_empty ivl2 then empty else*)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   135
     case (ivl1,ivl2) of (I l1 h1, I l2 h2) \<Rightarrow>
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   136
       I (if l1 = None then l2 else l1)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   137
         (if h1 = None then h2 else h1))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   138
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   139
instance
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   140
proof qed
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   141
  (auto simp add: widen_ivl_def narrow_ivl_def le_option_def le_ivl_def empty_def split: ivl.split option.split if_splits)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   142
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   143
end
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   144
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   145
instantiation astate :: (WN)WN
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   146
begin
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   147
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   148
definition "widen_astate F1 F2 =
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   149
  FunDom (\<lambda>x. fun F1 x \<nabla> fun F2 x) (inter_list (dom F1) (dom F2))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   150
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   151
definition "narrow_astate F1 F2 =
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   152
  FunDom (\<lambda>x. fun F1 x \<triangle> fun F2 x) (inter_list (dom F1) (dom F2))"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   153
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   154
instance
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   155
proof
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   156
  case goal1 thus ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   157
    by(simp add: widen_astate_def le_astate_def lookup_def widen)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   158
next
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   159
  case goal2 thus ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   160
    by(auto simp: narrow_astate_def le_astate_def lookup_def narrow1)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   161
next
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   162
  case goal3 thus ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   163
    by(auto simp: narrow_astate_def le_astate_def lookup_def narrow2)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   164
qed
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   165
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   166
end
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   167
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   168
instantiation up :: (WN)WN
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   169
begin
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   170
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   171
fun widen_up where
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   172
"widen_up bot x = x" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   173
"widen_up x bot = x" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   174
"widen_up (Up x) (Up y) = Up(x \<nabla> y)"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   175
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   176
fun narrow_up where
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   177
"narrow_up bot x = bot" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   178
"narrow_up x bot = bot" |
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   179
"narrow_up (Up x) (Up y) = Up(x \<triangle> y)"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   180
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   181
instance
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   182
proof
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   183
  case goal1 show ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   184
    by(induct x y rule: widen_up.induct) (simp_all add: widen)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   185
next
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   186
  case goal2 thus ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   187
    by(induct x y rule: narrow_up.induct) (simp_all add: narrow1)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   188
next
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   189
  case goal3 thus ?case
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   190
    by(induct x y rule: narrow_up.induct) (simp_all add: narrow2)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   191
qed
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   192
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   193
end
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   194
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   195
interpretation
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
   196
  Abs_Int1 rep_ivl num_ivl plus_ivl inv_plus_ivl inv_less_ivl "(iter' 3 2)"
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   197
defines afilter_ivl' is afilter
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   198
and bfilter_ivl' is bfilter
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   199
and AI_ivl' is AI
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   200
and aval_ivl' is aval'
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
   201
proof qed (auto simp: iter'_pfp_above)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   202
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   203
value [code] "list_up(AI_ivl' test3_ivl Top)"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   204
value [code] "list_up(AI_ivl' test4_ivl Top)"
45020
21334181f820 added example
nipkow
parents: 45019
diff changeset
   205
value [code] "list_up(AI_ivl' test5_ivl Top)"
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   206
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents:
diff changeset
   207
end