src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy
author haftmann
Sun, 21 Sep 2014 16:56:11 +0200
changeset 58410 6d46ad54a2ab
parent 58310 91ea607a34d8
child 61671 20d4cd2ceab2
permissions -rw-r--r--
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     2
45111
054a9ac0d7ef Added Hoare-like Abstract Interpretation
nipkow
parents: 45110
diff changeset
     3
theory Abs_Int_den1_ivl
45990
b7b905b23b2a incorporated More_Set and More_List into the Main body -- to be consolidated later
haftmann
parents: 45978
diff changeset
     4
imports Abs_Int_den1
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     5
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     6
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     7
subsection "Interval Analysis"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     8
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
     9
datatype ivl = I "int option" "int option"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    10
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    11
text{* We assume an important invariant: arithmetic operations are never
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    12
applied to empty intervals @{term"I (Some i) (Some j)"} with @{term"j <
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    13
i"}. This avoids special cases. Why can we assume this? Because an empty
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    14
interval of values for a variable means that the current program point is
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    15
unreachable. But this should actually translate into the bottom state, not a
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    16
state where some variables have empty intervals. *}
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    17
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    18
definition "rep_ivl i =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    19
 (case i of I (Some l) (Some h) \<Rightarrow> {l..h} | I (Some l) None \<Rightarrow> {l..}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    20
  | I None (Some h) \<Rightarrow> {..h} | I None None \<Rightarrow> UNIV)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    21
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    22
definition "num_ivl n = I (Some n) (Some n)"
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    23
45978
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    24
definition
46028
9f113cdf3d66 attribute code_abbrev superseedes code_unfold_post
haftmann
parents: 45990
diff changeset
    25
  [code_abbrev]: "contained_in i k \<longleftrightarrow> k \<in> rep_ivl i"
45978
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    26
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    27
lemma contained_in_simps [code]:
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    28
  "contained_in (I (Some l) (Some h)) k \<longleftrightarrow> l \<le> k \<and> k \<le> h"
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    29
  "contained_in (I (Some l) None) k \<longleftrightarrow> l \<le> k"
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    30
  "contained_in (I None (Some h)) k \<longleftrightarrow> k \<le> h"
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    31
  "contained_in (I None None) k \<longleftrightarrow> True"
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    32
  by (simp_all add: contained_in_def rep_ivl_def)
d3325de5f299 executable intervals
haftmann
parents: 45111
diff changeset
    33
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    34
instantiation option :: (plus)plus
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    35
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    36
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    37
fun plus_option where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    38
"Some x + Some y = Some(x+y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    39
"_ + _ = None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    40
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    41
instance proof qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    42
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    43
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    44
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    45
definition empty where "empty = I (Some 1) (Some 0)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    46
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    47
fun is_empty where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    48
"is_empty(I (Some l) (Some h)) = (h<l)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    49
"is_empty _ = False"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    50
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    51
lemma [simp]: "is_empty(I l h) =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    52
  (case l of Some l \<Rightarrow> (case h of Some h \<Rightarrow> h<l | None \<Rightarrow> False) | None \<Rightarrow> False)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    53
by(auto split:option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    54
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    55
lemma [simp]: "is_empty i \<Longrightarrow> rep_ivl i = {}"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    56
by(auto simp add: rep_ivl_def split: ivl.split option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    57
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    58
definition "plus_ivl i1 i2 = ((*if is_empty i1 | is_empty i2 then empty else*)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    59
  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1+l2) (h1+h2))"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    60
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    61
instantiation ivl :: SL_top
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    62
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    63
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    64
definition le_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> bool" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    65
"le_option pos x y =
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    66
 (case x of (Some i) \<Rightarrow> (case y of Some j \<Rightarrow> i\<le>j | None \<Rightarrow> pos)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
    67
  | None \<Rightarrow> (case y of Some j \<Rightarrow> \<not>pos | None \<Rightarrow> True))"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    68
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    69
fun le_aux where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    70
"le_aux (I l1 h1) (I l2 h2) = (le_option False l2 l1 & le_option True h1 h2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    71
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    72
definition le_ivl where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    73
"i1 \<sqsubseteq> i2 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    74
 (if is_empty i1 then True else
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    75
  if is_empty i2 then False else le_aux i1 i2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    76
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    77
definition min_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    78
"min_option pos o1 o2 = (if le_option pos o1 o2 then o1 else o2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    79
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    80
definition max_option :: "bool \<Rightarrow> int option \<Rightarrow> int option \<Rightarrow> int option" where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    81
"max_option pos o1 o2 = (if le_option pos o1 o2 then o2 else o1)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    82
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    83
definition "i1 \<squnion> i2 =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    84
 (if is_empty i1 then i2 else if is_empty i2 then i1
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    85
  else case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    86
          I (min_option False l1 l2) (max_option True h1 h2))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    87
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    88
definition "Top = I None None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    89
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    90
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    91
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    92
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    93
    by(cases x, simp add: le_ivl_def le_option_def split: option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    94
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    95
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    96
    by(cases x, cases y, cases z, auto simp: le_ivl_def le_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    97
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    98
  case goal3 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    99
    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   100
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   101
  case goal4 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   102
    by(cases x, cases y, simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   103
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   104
  case goal5 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   105
    by(cases x, cases y, cases z, auto simp add: le_ivl_def join_ivl_def le_option_def min_option_def max_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   106
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   107
  case goal6 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   108
    by(cases x, simp add: Top_ivl_def le_ivl_def le_option_def split: option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   109
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   110
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   111
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   112
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   113
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   114
instantiation ivl :: L_top_bot
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   115
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   116
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   117
definition "i1 \<sqinter> i2 = (if is_empty i1 \<or> is_empty i2 then empty else
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   118
  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow>
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   119
    I (max_option False l1 l2) (min_option True h1 h2))"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   120
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   121
definition "Bot = empty"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   122
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   123
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   124
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   125
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   126
    by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   127
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   128
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   129
    by (simp add:meet_ivl_def empty_def meet_ivl_def le_ivl_def le_option_def max_option_def min_option_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   130
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   131
  case goal3 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   132
    by (cases x, cases y, cases z, auto simp add: le_ivl_def meet_ivl_def empty_def le_option_def max_option_def min_option_def split: option.splits if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   133
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   134
  case goal4 show ?case by(cases x, simp add: Bot_ivl_def empty_def le_ivl_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   135
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   136
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   137
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   138
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   139
instantiation option :: (minus)minus
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   140
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   141
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   142
fun minus_option where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   143
"Some x - Some y = Some(x-y)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   144
"_ - _ = None"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   145
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   146
instance proof qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   147
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   148
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   149
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   150
definition "minus_ivl i1 i2 = ((*if is_empty i1 | is_empty i2 then empty else*)
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   151
  case (i1,i2) of (I l1 h1, I l2 h2) \<Rightarrow> I (l1-h2) (h1-l2))"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   152
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   153
lemma rep_minus_ivl:
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   154
  "n1 : rep_ivl i1 \<Longrightarrow> n2 : rep_ivl i2 \<Longrightarrow> n1-n2 : rep_ivl(minus_ivl i1 i2)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   155
by(auto simp add: minus_ivl_def rep_ivl_def split: ivl.splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   156
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   157
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   158
definition "filter_plus_ivl i i1 i2 = ((*if is_empty i then empty else*)
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   159
  i1 \<sqinter> minus_ivl i i2, i2 \<sqinter> minus_ivl i i1)"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   160
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   161
fun filter_less_ivl :: "bool \<Rightarrow> ivl \<Rightarrow> ivl \<Rightarrow> ivl * ivl" where
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   162
"filter_less_ivl res (I l1 h1) (I l2 h2) =
45019
nipkow
parents: 44944
diff changeset
   163
  ((*if is_empty(I l1 h1) \<or> is_empty(I l2 h2) then (empty, empty) else*)
nipkow
parents: 44944
diff changeset
   164
   if res
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   165
   then (I l1 (min_option True h1 (h2 - Some 1)),
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   166
         I (max_option False (l1 + Some 1) l2) h2)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   167
   else (I (max_option False l1 l2) h1, I l2 (min_option True h1 h2)))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   168
55599
6535c537b243 aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
haftmann
parents: 52046
diff changeset
   169
permanent_interpretation Rep rep_ivl
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   170
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   171
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   172
    by(auto simp: rep_ivl_def le_ivl_def le_option_def split: ivl.split option.split if_splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   173
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   174
55599
6535c537b243 aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
haftmann
parents: 52046
diff changeset
   175
permanent_interpretation Val_abs rep_ivl num_ivl plus_ivl
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   176
proof
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   177
  case goal1 thus ?case by(simp add: rep_ivl_def num_ivl_def)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   178
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   179
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   180
    by(auto simp add: rep_ivl_def plus_ivl_def split: ivl.split option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   181
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   182
55599
6535c537b243 aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
haftmann
parents: 52046
diff changeset
   183
permanent_interpretation Rep1 rep_ivl
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   184
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   185
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   186
    by(auto simp add: rep_ivl_def meet_ivl_def empty_def min_option_def max_option_def split: ivl.split option.split)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   187
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   188
  case goal2 show ?case by(auto simp add: Bot_ivl_def rep_ivl_def empty_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   189
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   190
55599
6535c537b243 aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
haftmann
parents: 52046
diff changeset
   191
permanent_interpretation
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   192
  Val_abs1 rep_ivl num_ivl plus_ivl filter_plus_ivl filter_less_ivl
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   193
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   194
  case goal1 thus ?case
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   195
    by(auto simp add: filter_plus_ivl_def)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56927
diff changeset
   196
      (metis rep_minus_ivl add_diff_cancel add.commute)+
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   197
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   198
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   199
    by(cases a1, cases a2,
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   200
      auto simp: rep_ivl_def min_option_def max_option_def le_option_def split: if_splits option.splits)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   201
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   202
55599
6535c537b243 aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
haftmann
parents: 52046
diff changeset
   203
permanent_interpretation
45023
76abd26e2e2d renamed inv -> filter
nipkow
parents: 45020
diff changeset
   204
  Abs_Int1 rep_ivl num_ivl plus_ivl filter_plus_ivl filter_less_ivl "(iter' 3)"
55600
3c7610b8dcfc more convenient syntax for permanent interpretation
haftmann
parents: 55599
diff changeset
   205
defining afilter_ivl = afilter
3c7610b8dcfc more convenient syntax for permanent interpretation
haftmann
parents: 55599
diff changeset
   206
and bfilter_ivl = bfilter
3c7610b8dcfc more convenient syntax for permanent interpretation
haftmann
parents: 55599
diff changeset
   207
and AI_ivl = AI
3c7610b8dcfc more convenient syntax for permanent interpretation
haftmann
parents: 55599
diff changeset
   208
and aval_ivl = aval'
44944
f136409c2cef tuned post fixpoint setup
nipkow
parents: 44932
diff changeset
   209
proof qed (auto simp: iter'_pfp_above)
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   210
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   211
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   212
fun list_up where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   213
"list_up bot = bot" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   214
"list_up (Up x) = Up(list x)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   215
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   216
value "list_up(afilter_ivl (N 5) (I (Some 4) (Some 5)) Top)"
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   217
value "list_up(afilter_ivl (N 5) (I (Some 4) (Some 4)) Top)"
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   218
value "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 4))
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   219
 (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   220
value "list_up(afilter_ivl (V ''x'') (I (Some 4) (Some 5))
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   221
 (Up(FunDom(Top(''x'':=I (Some 5) (Some 6))) [''x''])))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   222
value "list_up(afilter_ivl (Plus (V ''x'') (V ''x'')) (I (Some 0) (Some 10))
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   223
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   224
value "list_up(afilter_ivl (Plus (V ''x'') (N 7)) (I (Some 0) (Some 10))
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   225
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 100)))[''x''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   226
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   227
value "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   228
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 0)))[''x''])))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   229
value "list_up(bfilter_ivl (Less (V ''x'') (V ''x'')) True
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   230
  (Up(FunDom(Top(''x'':= I (Some 0) (Some 2)))[''x''])))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   231
value "list_up(bfilter_ivl (Less (V ''x'') (Plus (N 10) (V ''y''))) True
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   232
  (Up(FunDom(Top(''x'':= I (Some 15) (Some 20),''y'':= I (Some 5) (Some 7)))[''x'', ''y''])))"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   233
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   234
definition "test_ivl1 =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   235
 ''y'' ::= N 7;;
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   236
 IF Less (V ''x'') (V ''y'')
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   237
 THEN ''y'' ::= Plus (V ''y'') (V ''x'')
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   238
 ELSE ''x'' ::= Plus (V ''x'') (V ''y'')"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   239
value "list_up(AI_ivl test_ivl1 Top)"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   240
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   241
value "list_up (AI_ivl test3_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   242
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   243
value "list_up (AI_ivl test5_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   244
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   245
value "list_up (AI_ivl test6_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   246
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   247
definition "test2_ivl =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   248
 ''y'' ::= N 7;;
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   249
 WHILE Less (V ''x'') (N 100) DO ''y'' ::= Plus (V ''y'') (N 1)"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   250
value "list_up(AI_ivl test2_ivl Top)"
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   251
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   252
definition "test3_ivl =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   253
 ''x'' ::= N 0;; ''y'' ::= N 100;; ''z'' ::= Plus (V ''x'') (V ''y'');;
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   254
 WHILE Less (V ''x'') (N 11)
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
   255
 DO (''x'' ::= Plus (V ''x'') (N 1);; ''y'' ::= Plus (V ''y'') (N (- 1)))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   256
value "list_up(AI_ivl test3_ivl Top)"
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   257
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   258
definition "test4_ivl =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   259
 ''x'' ::= N 0;; ''y'' ::= N 0;;
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   260
 WHILE Less (V ''x'') (N 1001)
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   261
 DO (''y'' ::= V ''x'';; ''x'' ::= Plus (V ''x'') (N 1))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   262
value "list_up(AI_ivl test4_ivl Top)"
44932
7c93ee993cae revised AbsInt and added widening and narrowing
nipkow
parents: 44923
diff changeset
   263
45020
21334181f820 added example
nipkow
parents: 45019
diff changeset
   264
text{* Nontermination not detected: *}
21334181f820 added example
nipkow
parents: 45019
diff changeset
   265
definition "test5_ivl =
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 46028
diff changeset
   266
 ''x'' ::= N 0;;
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58310
diff changeset
   267
 WHILE Less (V ''x'') (N 1) DO ''x'' ::= Plus (V ''x'') (N (- 1))"
56927
4044a7d1720f hardcoded nbe and sml into value command
haftmann
parents: 55600
diff changeset
   268
value "list_up(AI_ivl test5_ivl Top)"
45020
21334181f820 added example
nipkow
parents: 45019
diff changeset
   269
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   270
end