src/HOL/IMP/Abs_Int3.thy
author wenzelm
Sat, 09 Mar 2024 11:05:32 +0100
changeset 79823 60f1e32792c1
parent 73932 fd21b4a93043
child 80914 d97fdabd9e2b
permissions -rw-r--r--
tuned signature: prefer bottom-up construction;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     2
68778
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     3
subsection "Widening and Narrowing"
4566bac4517d improved sectioning
nipkow
parents: 67613
diff changeset
     4
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     5
theory Abs_Int3
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     6
imports Abs_Int2_ivl
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     7
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     8
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     9
class widen =
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    10
fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    11
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    12
class narrow =
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    13
fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    14
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
    15
class wn = widen + narrow + order +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    16
assumes widen1: "x \<le> x \<nabla> y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    17
assumes widen2: "y \<le> x \<nabla> y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    18
assumes narrow1: "y \<le> x \<Longrightarrow> y \<le> x \<triangle> y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    19
assumes narrow2: "y \<le> x \<Longrightarrow> x \<triangle> y \<le> x"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    20
begin
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    21
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    22
lemma narrowid[simp]: "x \<triangle> x = x"
73411
1f1366966296 avoid name clash
haftmann
parents: 69597
diff changeset
    23
by (rule order.antisym) (simp_all add: narrow1 narrow2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    24
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    25
end
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    26
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52504
diff changeset
    27
lemma top_widen_top[simp]: "\<top> \<nabla> \<top> = (\<top>::_::{wn,order_top})"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    28
by (metis eq_iff top_greatest widen2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    29
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
    30
instantiation ivl :: wn
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    31
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    32
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    33
definition "widen_rep p1 p2 =
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    34
  (if is_empty_rep p1 then p2 else if is_empty_rep p2 then p1 else
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    35
   let (l1,h1) = p1; (l2,h2) = p2
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    36
   in (if l2 < l1 then Minf else l1, if h1 < h2 then Pinf else h1))"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    37
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    38
lift_definition widen_ivl :: "ivl \<Rightarrow> ivl \<Rightarrow> ivl" is widen_rep
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    39
by(auto simp: widen_rep_def eq_ivl_iff)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    40
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    41
definition "narrow_rep p1 p2 =
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    42
  (if is_empty_rep p1 \<or> is_empty_rep p2 then empty_rep else
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    43
   let (l1,h1) = p1; (l2,h2) = p2
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    44
   in (if l1 = Minf then l2 else l1, if h1 = Pinf then h2 else h1))"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    45
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    46
lift_definition narrow_ivl :: "ivl \<Rightarrow> ivl \<Rightarrow> ivl" is narrow_rep
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    47
by(auto simp: narrow_rep_def eq_ivl_iff)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    48
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    49
instance
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
    50
proof
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    51
qed (transfer, auto simp: widen_rep_def narrow_rep_def le_iff_subset \<gamma>_rep_def subset_eq is_empty_rep_def empty_rep_def eq_ivl_def split: if_splits extended.splits)+
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    52
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    53
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    54
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52504
diff changeset
    55
instantiation st :: ("{order_top,wn}")wn
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    56
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    57
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
    58
lift_definition widen_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (\<nabla>)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    59
by(auto simp: eq_st_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    60
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
    61
lift_definition narrow_st :: "'a st \<Rightarrow> 'a st \<Rightarrow> 'a st" is "map2_st_rep (\<triangle>)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
    62
by(auto simp: eq_st_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    63
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    64
instance
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    65
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    66
  case 1 thus ?case by transfer (simp add: less_eq_st_rep_iff widen1)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    67
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    68
  case 2 thus ?case by transfer (simp add: less_eq_st_rep_iff widen2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    69
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    70
  case 3 thus ?case by transfer (simp add: less_eq_st_rep_iff narrow1)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    71
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    72
  case 4 thus ?case by transfer (simp add: less_eq_st_rep_iff narrow2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    73
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    74
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    75
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    76
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    77
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
    78
instantiation option :: (wn)wn
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    79
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    80
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    81
fun widen_option where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    82
"None \<nabla> x = x" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    83
"x \<nabla> None = x" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    84
"(Some x) \<nabla> (Some y) = Some(x \<nabla> y)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    85
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    86
fun narrow_option where
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    87
"None \<triangle> x = None" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    88
"x \<triangle> None = None" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    89
"(Some x) \<triangle> (Some y) = Some(x \<triangle> y)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    90
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    91
instance
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    92
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    93
  case (1 x y) thus ?case
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    94
    by(induct x y rule: widen_option.induct)(simp_all add: widen1)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    95
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    96
  case (2 x y) thus ?case
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    97
    by(induct x y rule: widen_option.induct)(simp_all add: widen2)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    98
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
    99
  case (3 x y) thus ?case
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   100
    by(induct x y rule: narrow_option.induct) (simp_all add: narrow1)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   101
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   102
  case (4 y x) thus ?case
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   103
    by(induct x y rule: narrow_option.induct) (simp_all add: narrow2)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   104
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   105
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   106
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   107
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   108
definition map2_acom :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a acom \<Rightarrow> 'a acom \<Rightarrow> 'a acom"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   109
where
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   110
"map2_acom f C1 C2 = annotate (\<lambda>p. f (anno C1 p) (anno C2 p)) (strip C1)"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   111
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   112
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   113
instantiation acom :: (widen)widen
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   114
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
   115
definition "widen_acom = map2_acom (\<nabla>)"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   116
instance ..
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   117
end
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   118
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   119
instantiation acom :: (narrow)narrow
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   120
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
   121
definition "narrow_acom = map2_acom (\<triangle>)"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   122
instance ..
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   123
end
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   124
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   125
lemma strip_map2_acom[simp]:
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   126
 "strip C1 = strip C2 \<Longrightarrow> strip(map2_acom f C1 C2) = strip C1"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   127
by(simp add: map2_acom_def)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   128
(*by(induct f C1 C2 rule: map2_acom.induct) simp_all*)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   129
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   130
lemma strip_widen_acom[simp]:
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   131
  "strip C1 = strip C2 \<Longrightarrow> strip(C1 \<nabla> C2) = strip C1"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   132
by(simp add: widen_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   133
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   134
lemma strip_narrow_acom[simp]:
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   135
  "strip C1 = strip C2 \<Longrightarrow> strip(C1 \<triangle> C2) = strip C1"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   136
by(simp add: narrow_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   137
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   138
lemma narrow1_acom: "C2 \<le> C1 \<Longrightarrow> C2 \<le> C1 \<triangle> (C2::'a::wn acom)"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   139
by(simp add: narrow_acom_def narrow1 map2_acom_def less_eq_acom_def size_annos)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   140
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   141
lemma narrow2_acom: "C2 \<le> C1 \<Longrightarrow> C1 \<triangle> (C2::'a::wn acom) \<le> C1"
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   142
by(simp add: narrow_acom_def narrow2 map2_acom_def less_eq_acom_def size_annos)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   143
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   144
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   145
subsubsection "Pre-fixpoint computation"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   146
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   147
definition iter_widen :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> ('a::{order,widen})option"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   148
where "iter_widen f = while_option (\<lambda>x. \<not> f x \<le> x) (\<lambda>x. x \<nabla> f x)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   149
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   150
definition iter_narrow :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> ('a::{order,narrow})option"
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   151
where "iter_narrow f = while_option (\<lambda>x. x \<triangle> f x < x) (\<lambda>x. x \<triangle> f x)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   152
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   153
definition pfp_wn :: "('a::{order,widen,narrow} \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   154
where "pfp_wn f x =
49576
nipkow
parents: 49548
diff changeset
   155
  (case iter_widen f x of None \<Rightarrow> None | Some p \<Rightarrow> iter_narrow f p)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   156
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   157
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   158
lemma iter_widen_pfp: "iter_widen f x = Some p \<Longrightarrow> f p \<le> p"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   159
by(auto simp add: iter_widen_def dest: while_option_stop)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   160
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   161
lemma iter_widen_inv:
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   162
assumes "!!x. P x \<Longrightarrow> P(f x)" "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<nabla> x2)" and "P x"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   163
and "iter_widen f x = Some y" shows "P y"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   164
using while_option_rule[where P = "P", OF _ assms(4)[unfolded iter_widen_def]]
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   165
by (blast intro: assms(1-3))
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   166
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   167
lemma strip_while: fixes f :: "'a acom \<Rightarrow> 'a acom"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   168
assumes "\<forall>C. strip (f C) = strip C" and "while_option P f C = Some C'"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   169
shows "strip C' = strip C"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   170
using while_option_rule[where P = "\<lambda>C'. strip C' = strip C", OF _ assms(2)]
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   171
by (metis assms(1))
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   172
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   173
lemma strip_iter_widen: fixes f :: "'a::{order,widen} acom \<Rightarrow> 'a acom"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   174
assumes "\<forall>C. strip (f C) = strip C" and "iter_widen f C = Some C'"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   175
shows "strip C' = strip C"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   176
proof-
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   177
  have "\<forall>C. strip(C \<nabla> f C) = strip C"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   178
    by (metis assms(1) strip_map2_acom widen_acom_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   179
  from strip_while[OF this] assms(2) show ?thesis by(simp add: iter_widen_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   180
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   181
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   182
lemma iter_narrow_pfp:
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   183
assumes mono: "!!x1 x2::_::wn acom. P x1 \<Longrightarrow> P x2 \<Longrightarrow> x1 \<le> x2 \<Longrightarrow> f x1 \<le> f x2"
49576
nipkow
parents: 49548
diff changeset
   184
and Pinv: "!!x. P x \<Longrightarrow> P(f x)" "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<triangle> x2)"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   185
and "P p0" and "f p0 \<le> p0" and "iter_narrow f p0 = Some p"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   186
shows "P p \<and> f p \<le> p"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   187
proof-
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   188
  let ?Q = "%p. P p \<and> f p \<le> p \<and> p \<le> p0"
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   189
  have "?Q (p \<triangle> f p)" if Q: "?Q p" for p
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   190
  proof auto
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   191
    note P = conjunct1[OF Q] and 12 = conjunct2[OF Q]
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   192
    note 1 = conjunct1[OF 12] and 2 = conjunct2[OF 12]
49576
nipkow
parents: 49548
diff changeset
   193
    let ?p' = "p \<triangle> f p"
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   194
    show "P ?p'" by (blast intro: P Pinv)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   195
    have "f ?p' \<le> f p" by(rule mono[OF \<open>P (p \<triangle> f p)\<close>  P narrow2_acom[OF 1]])
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   196
    also have "\<dots> \<le> ?p'" by(rule narrow1_acom[OF 1])
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   197
    finally show "f ?p' \<le> ?p'" .
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   198
    have "?p' \<le> p" by (rule narrow2_acom[OF 1])
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   199
    also have "p \<le> p0" by(rule 2)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   200
    finally show "?p' \<le> p0" .
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 64267
diff changeset
   201
  qed
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   202
  thus ?thesis
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   203
    using while_option_rule[where P = ?Q, OF _ assms(6)[simplified iter_narrow_def]]
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   204
    by (blast intro: assms(4,5) le_refl)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   205
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   206
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   207
lemma pfp_wn_pfp:
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   208
assumes mono: "!!x1 x2::_::wn acom. P x1 \<Longrightarrow> P x2 \<Longrightarrow> x1 \<le> x2 \<Longrightarrow> f x1 \<le> f x2"
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   209
and Pinv: "P x"  "!!x. P x \<Longrightarrow> P(f x)"
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   210
  "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<nabla> x2)"
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   211
  "!!x1 x2. P x1 \<Longrightarrow> P x2 \<Longrightarrow> P(x1 \<triangle> x2)"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   212
and pfp_wn: "pfp_wn f x = Some p" shows "P p \<and> f p \<le> p"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   213
proof-
49576
nipkow
parents: 49548
diff changeset
   214
  from pfp_wn obtain p0
nipkow
parents: 49548
diff changeset
   215
    where its: "iter_widen f x = Some p0" "iter_narrow f p0 = Some p"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   216
    by(auto simp: pfp_wn_def split: option.splits)
49576
nipkow
parents: 49548
diff changeset
   217
  have "P p0" by (blast intro: iter_widen_inv[where P="P"] its(1) Pinv(1-3))
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   218
  thus ?thesis
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   219
    by - (assumption |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   220
          rule iter_narrow_pfp[where P=P] mono Pinv(2,4) iter_widen_pfp its)+
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   221
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   222
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   223
lemma strip_pfp_wn:
49548
8192dc55bda9 generalized types
nipkow
parents: 49547
diff changeset
   224
  "\<lbrakk> \<forall>C. strip(f C) = strip C; pfp_wn f C = Some C' \<rbrakk> \<Longrightarrow> strip C' = strip C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   225
by(auto simp add: pfp_wn_def iter_narrow_def split: option.splits)
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51385
diff changeset
   226
  (metis (mono_tags) strip_iter_widen strip_narrow_acom strip_while)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   227
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   228
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   229
locale Abs_Int_wn = Abs_Int_inv_mono where \<gamma>=\<gamma>
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   230
  for \<gamma> :: "'av::{wn,bounded_lattice} \<Rightarrow> val set"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   231
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   232
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   233
definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   234
"AI_wn c = pfp_wn (step' \<top>) (bot c)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   235
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   236
lemma AI_wn_correct: "AI_wn c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   237
proof(simp add: CS_def AI_wn_def)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   238
  assume 1: "pfp_wn (step' \<top>) (bot c) = Some C"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   239
  have 2: "strip C = c \<and> step' \<top> C \<le> C"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   240
    by(rule pfp_wn_pfp[where x="bot c"]) (simp_all add: 1 mono_step'_top)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   241
  have pfp: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C"
50986
c54ea7f5418f simplified proofs
nipkow
parents: 49892
diff changeset
   242
  proof(rule order_trans)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   243
    show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le>  \<gamma>\<^sub>c (step' \<top> C)"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   244
      by(rule step_step')
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   245
    show "... \<le> \<gamma>\<^sub>c C"
50986
c54ea7f5418f simplified proofs
nipkow
parents: 49892
diff changeset
   246
      by(rule mono_gamma_c[OF conjunct2[OF 2]])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   247
  qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   248
  have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp_wn[OF _ 1])
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   249
  have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   250
    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 pfp])
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   251
  thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   252
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   253
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   254
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   255
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   256
global_interpretation Abs_Int_wn
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
   257
where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = "(+)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   258
and test_num' = in_ivl
51974
9c80e62161a5 tuned names
nipkow
parents: 51953
diff changeset
   259
and inv_plus' = inv_plus_ivl and inv_less' = inv_less_ivl
61671
20d4cd2ceab2 prefer "rewrites" and "defines" to note rewrite morphisms
haftmann
parents: 61179
diff changeset
   260
defines AI_wn_ivl = AI_wn
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   261
..
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   262
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   263
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   264
subsubsection "Tests"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   265
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   266
definition "step_up_ivl n = ((\<lambda>C. C \<nabla> step_ivl \<top> C)^^n)"
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   267
definition "step_down_ivl n = ((\<lambda>C. C \<triangle> step_ivl \<top> C)^^n)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   268
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   269
text\<open>For \<^const>\<open>test3_ivl\<close>, \<^const>\<open>AI_ivl\<close> needed as many iterations as
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   270
the loop took to execute. In contrast, \<^const>\<open>AI_wn_ivl\<close> converges in a
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   271
constant number of steps:\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   272
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   273
value "show_acom (step_up_ivl 1 (bot test3_ivl))"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   274
value "show_acom (step_up_ivl 2 (bot test3_ivl))"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   275
value "show_acom (step_up_ivl 3 (bot test3_ivl))"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   276
value "show_acom (step_up_ivl 4 (bot test3_ivl))"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   277
value "show_acom (step_up_ivl 5 (bot test3_ivl))"
49188
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   278
value "show_acom (step_up_ivl 6 (bot test3_ivl))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   279
value "show_acom (step_up_ivl 7 (bot test3_ivl))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   280
value "show_acom (step_up_ivl 8 (bot test3_ivl))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   281
value "show_acom (step_down_ivl 1 (step_up_ivl 8 (bot test3_ivl)))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   282
value "show_acom (step_down_ivl 2 (step_up_ivl 8 (bot test3_ivl)))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   283
value "show_acom (step_down_ivl 3 (step_up_ivl 8 (bot test3_ivl)))"
22f7e7b68f50 adjusted examples
nipkow
parents: 49095
diff changeset
   284
value "show_acom (step_down_ivl 4 (step_up_ivl 8 (bot test3_ivl)))"
51953
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   285
value "show_acom_opt (AI_wn_ivl test3_ivl)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   286
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   287
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   288
text\<open>Now all the analyses terminate:\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   289
51953
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   290
value "show_acom_opt (AI_wn_ivl test4_ivl)"
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   291
value "show_acom_opt (AI_wn_ivl test5_ivl)"
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   292
value "show_acom_opt (AI_wn_ivl test6_ivl)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   293
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   294
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   295
subsubsection "Generic Termination Proof"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   296
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51711
diff changeset
   297
lemma top_on_opt_widen:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   298
  "top_on_opt o1 X \<Longrightarrow> top_on_opt o2 X \<Longrightarrow> top_on_opt (o1 \<nabla> o2 :: _ st option) X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   299
apply(induct o1 o2 rule: widen_option.induct)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   300
apply (auto)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   301
by transfer simp
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   302
51722
3da99469cc1b proved termination for fun-based AI
nipkow
parents: 51711
diff changeset
   303
lemma top_on_opt_narrow:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   304
  "top_on_opt o1 X \<Longrightarrow> top_on_opt o2 X \<Longrightarrow> top_on_opt (o1 \<triangle> o2 :: _ st option) X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   305
apply(induct o1 o2 rule: narrow_option.induct)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   306
apply (auto)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   307
by transfer simp
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   308
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   309
(* FIXME mk anno abbrv *)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   310
lemma annos_map2_acom[simp]: "strip C2 = strip C1 \<Longrightarrow>
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   311
  annos(map2_acom f C1 C2) = map (%(x,y).f x y) (zip (annos C1) (annos C2))"
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   312
by(simp add: map2_acom_def list_eq_iff_nth_eq size_annos anno_def[symmetric] size_annos_same[of C1 C2])
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   313
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   314
lemma top_on_acom_widen:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   315
  "\<lbrakk>top_on_acom C1 X; strip C1 = strip C2; top_on_acom C2 X\<rbrakk>
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   316
  \<Longrightarrow> top_on_acom (C1 \<nabla> C2 :: _ st option acom) X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   317
by(auto simp add: widen_acom_def top_on_acom_def)(metis top_on_opt_widen in_set_zipE)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   318
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   319
lemma top_on_acom_narrow:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   320
  "\<lbrakk>top_on_acom C1 X; strip C1 = strip C2; top_on_acom C2 X\<rbrakk>
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   321
  \<Longrightarrow> top_on_acom (C1 \<triangle> C2 :: _ st option acom) X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   322
by(auto simp add: narrow_acom_def top_on_acom_def)(metis top_on_opt_narrow in_set_zipE)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   323
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   324
text\<open>The assumptions for widening and narrowing differ because during
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   325
narrowing we have the invariant \<^prop>\<open>y \<le> x\<close> (where \<open>y\<close> is the next
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   326
iterate), but during widening there is no such invariant, there we only have
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   327
that not yet \<^prop>\<open>y \<le> x\<close>. This complicates the termination proof for
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   328
widening.\<close>
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   329
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   330
locale Measure_wn = Measure1 where m=m
52729
412c9e0381a1 factored syntactic type classes for bot and top (by Alessandro Coglio)
haftmann
parents: 52504
diff changeset
   331
  for m :: "'av::{order_top,wn} \<Rightarrow> nat" +
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   332
fixes n :: "'av \<Rightarrow> nat"
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   333
assumes m_anti_mono: "x \<le> y \<Longrightarrow> m x \<ge> m y"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   334
assumes m_widen: "~ y \<le> x \<Longrightarrow> m(x \<nabla> y) < m x"
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   335
assumes n_narrow: "y \<le> x \<Longrightarrow> x \<triangle> y < x \<Longrightarrow> n(x \<triangle> y) < n x"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   336
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   337
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   338
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   339
lemma m_s_anti_mono_rep: assumes "\<forall>x. S1 x \<le> S2 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   340
shows "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   341
proof-
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   342
  from assms have "\<forall>x. m(S1 x) \<ge> m(S2 x)" by (metis m_anti_mono)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   343
  thus "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))" by (metis sum_mono)
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   344
qed
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   345
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   346
lemma m_s_anti_mono: "S1 \<le> S2 \<Longrightarrow> m_s S1 X \<ge> m_s S2 X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   347
unfolding m_s_def
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   348
apply (transfer fixing: m)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   349
apply(simp add: less_eq_st_rep_iff eq_st_def m_s_anti_mono_rep)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   350
done
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   351
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   352
lemma m_s_widen_rep: assumes "finite X" "S1 = S2 on -X" "\<not> S2 x \<le> S1 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   353
  shows "(\<Sum>x\<in>X. m (S1 x \<nabla> S2 x)) < (\<Sum>x\<in>X. m (S1 x))"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   354
proof-
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   355
  have 1: "\<forall>x\<in>X. m(S1 x) \<ge> m(S1 x \<nabla> S2 x)"
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   356
    by (metis m_anti_mono wn_class.widen1)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   357
  have "x \<in> X" using assms(2,3)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   358
    by(auto simp add: Ball_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   359
  hence 2: "\<exists>x\<in>X. m(S1 x) > m(S1 x \<nabla> S2 x)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   360
    using assms(3) m_widen by blast
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   361
  from sum_strict_mono_ex1[OF \<open>finite X\<close> 1 2]
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   362
  show ?thesis .
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   363
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   364
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   365
lemma m_s_widen: "finite X \<Longrightarrow> fun S1 = fun S2 on -X ==>
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   366
  ~ S2 \<le> S1 \<Longrightarrow> m_s (S1 \<nabla> S2) X < m_s S1 X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   367
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
   368
apply (transfer fixing: m)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   369
apply(auto simp add: less_eq_st_rep_iff m_s_widen_rep)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   370
done
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   371
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   372
lemma m_o_anti_mono: "finite X \<Longrightarrow> top_on_opt o1 (-X) \<Longrightarrow> top_on_opt o2 (-X) \<Longrightarrow>
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   373
  o1 \<le> o2 \<Longrightarrow> m_o o1 X \<ge> m_o o2 X"
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   374
proof(induction o1 o2 rule: less_eq_option.induct)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   375
  case 1 thus ?case by (simp add: m_o_def)(metis m_s_anti_mono)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   376
next
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   377
  case 2 thus ?case
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   378
    by(simp add: m_o_def le_SucI m_s_h split: option.splits)
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   379
next
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   380
  case 3 thus ?case by simp
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   381
qed
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   382
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   383
lemma m_o_widen: "\<lbrakk> finite X; top_on_opt S1 (-X); top_on_opt S2 (-X); \<not> S2 \<le> S1 \<rbrakk> \<Longrightarrow>
51791
c4db685eaed0 more standard argument order
nipkow
parents: 51786
diff changeset
   384
  m_o (S1 \<nabla> S2) X < m_o S1 X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   385
by(auto simp: m_o_def m_s_h less_Suc_eq_le m_s_widen split: option.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   386
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   387
lemma m_c_widen:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   388
  "strip C1 = strip C2  \<Longrightarrow> top_on_acom C1 (-vars C1) \<Longrightarrow> top_on_acom C2 (-vars C2)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   389
   \<Longrightarrow> \<not> C2 \<le> C1 \<Longrightarrow> m_c (C1 \<nabla> C2) < m_c C1"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   390
apply(auto simp: m_c_def widen_acom_def map2_acom_def size_annos[symmetric] anno_def[symmetric]sum_list_sum_nth)
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   391
apply(subgoal_tac "length(annos C2) = length(annos C1)")
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51385
diff changeset
   392
 prefer 2 apply (simp add: size_annos_same2)
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   393
apply (auto)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   394
apply(rule sum_strict_mono_ex1)
52019
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   395
 apply(auto simp add: m_o_anti_mono vars_acom_def anno_def top_on_acom_def top_on_opt_widen widen1 less_eq_acom_def listrel_iff_nth)
a4cbca8f7342 finally: acom with pointwise access and update of annotations
nipkow
parents: 51974
diff changeset
   396
apply(rule_tac x=p in bexI)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   397
 apply (auto simp: vars_acom_def m_o_widen top_on_acom_def)
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   398
done
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   399
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   400
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   401
definition n_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>s") where
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   402
"n\<^sub>s S X = (\<Sum>x\<in>X. n(fun S x))"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   403
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   404
lemma n_s_narrow_rep:
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   405
assumes "finite X"  "S1 = S2 on -X"  "\<forall>x. S2 x \<le> S1 x"  "\<forall>x. S1 x \<triangle> S2 x \<le> S1 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   406
  "S1 x \<noteq> S1 x \<triangle> S2 x"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   407
shows "(\<Sum>x\<in>X. n (S1 x \<triangle> S2 x)) < (\<Sum>x\<in>X. n (S1 x))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   408
proof-
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   409
  have 1: "\<forall>x. n(S1 x \<triangle> S2 x) \<le> n(S1 x)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   410
      by (metis assms(3) assms(4) eq_iff less_le_not_le n_narrow)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   411
  have "x \<in> X" by (metis Compl_iff assms(2) assms(5) narrowid)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   412
  hence 2: "\<exists>x\<in>X. n(S1 x \<triangle> S2 x) < n(S1 x)"
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   413
    by (metis assms(3-5) eq_iff less_le_not_le n_narrow)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   414
  show ?thesis
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   415
    apply(rule sum_strict_mono_ex1[OF \<open>finite X\<close>]) using 1 2 by blast+
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   416
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   417
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   418
lemma n_s_narrow: "finite X \<Longrightarrow> fun S1 = fun S2 on -X \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   419
  \<Longrightarrow> n\<^sub>s (S1 \<triangle> S2) X < n\<^sub>s S1 X"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   420
apply(auto simp add: less_st_def n_s_def)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   421
apply (transfer fixing: n)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   422
apply(auto simp add: less_eq_st_rep_iff eq_st_def fun_eq_iff n_s_narrow_rep)
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   423
done
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   424
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   425
definition n_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>o") where
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   426
"n\<^sub>o opt X = (case opt of None \<Rightarrow> 0 | Some S \<Rightarrow> n\<^sub>s S X + 1)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   427
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   428
lemma n_o_narrow:
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   429
  "top_on_opt S1 (-X) \<Longrightarrow> top_on_opt S2 (-X) \<Longrightarrow> finite X
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   430
  \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1 \<Longrightarrow> n\<^sub>o (S1 \<triangle> S2) X < n\<^sub>o S1 X"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   431
apply(induction S1 S2 rule: narrow_option.induct)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   432
apply(auto simp: n_o_def n_s_narrow)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   433
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   434
49576
nipkow
parents: 49548
diff changeset
   435
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   436
definition n_c :: "'av st option acom \<Rightarrow> nat" ("n\<^sub>c") where
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 61890
diff changeset
   437
"n\<^sub>c C = sum_list (map (\<lambda>a. n\<^sub>o a (vars C)) (annos C))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   438
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   439
lemma less_annos_iff: "(C1 < C2) = (C1 \<le> C2 \<and>
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   440
  (\<exists>i<length (annos C1). annos C1 ! i < annos C2 ! i))"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73411
diff changeset
   441
by(metis (opaque_lifting, no_types) less_le_not_le le_iff_le_annos size_annos_same2)
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   442
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   443
lemma n_c_narrow: "strip C1 = strip C2
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   444
  \<Longrightarrow> top_on_acom C1 (- vars C1) \<Longrightarrow> top_on_acom C2 (- vars C2)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52729
diff changeset
   445
  \<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n\<^sub>c (C1 \<triangle> C2) < n\<^sub>c C1"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   446
apply(auto simp: n_c_def narrow_acom_def sum_list_sum_nth)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   447
apply(subgoal_tac "length(annos C2) = length(annos C1)")
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   448
prefer 2 apply (simp add: size_annos_same2)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   449
apply (auto)
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   450
apply(simp add: less_annos_iff le_iff_le_annos)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63882
diff changeset
   451
apply(rule sum_strict_mono_ex1)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   452
apply (auto simp: vars_acom_def top_on_acom_def)
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   453
apply (metis n_o_narrow nth_mem finite_cvars less_imp_le le_less order_refl)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   454
apply(rule_tac x=i in bexI)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   455
prefer 2 apply simp
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   456
apply(rule n_o_narrow[where X = "vars(strip C2)"])
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   457
apply (simp_all)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   458
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   459
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   460
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   461
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   462
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   463
lemma iter_widen_termination:
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   464
fixes m :: "'a::wn acom \<Rightarrow> nat"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   465
assumes P_f: "\<And>C. P C \<Longrightarrow> P(f C)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   466
and P_widen: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> P(C1 \<nabla> C2)"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   467
and m_widen: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> ~ C2 \<le> C1 \<Longrightarrow> m(C1 \<nabla> C2) < m C1"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   468
and "P C" shows "\<exists>C'. iter_widen f C = Some C'"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   469
proof(simp add: iter_widen_def,
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   470
      rule measure_while_option_Some[where P = P and f=m])
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   471
  show "P C" by(rule \<open>P C\<close>)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   472
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   473
  fix C assume "P C" "\<not> f C \<le> C" thus "P (C \<nabla> f C) \<and> m (C \<nabla> f C) < m C"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   474
    by(simp add: P_f P_widen m_widen)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   475
qed
49496
2694d1615eef more termination proofs
nipkow
parents: 49399
diff changeset
   476
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   477
lemma iter_narrow_termination:
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   478
fixes n :: "'a::wn acom \<Rightarrow> nat"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   479
assumes P_f: "\<And>C. P C \<Longrightarrow> P(f C)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   480
and P_narrow: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> P(C1 \<triangle> C2)"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   481
and mono: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> f C1 \<le> f C2"
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   482
and n_narrow: "\<And>C1 C2. P C1 \<Longrightarrow> P C2 \<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n(C1 \<triangle> C2) < n C1"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   483
and init: "P C" "f C \<le> C" shows "\<exists>C'. iter_narrow f C = Some C'"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   484
proof(simp add: iter_narrow_def,
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   485
      rule measure_while_option_Some[where f=n and P = "%C. P C \<and> f C \<le> C"])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   486
  show "P C \<and> f C \<le> C" using init by blast
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   487
next
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   488
  fix C assume 1: "P C \<and> f C \<le> C" and 2: "C \<triangle> f C < C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   489
  hence "P (C \<triangle> f C)" by(simp add: P_f P_narrow)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   490
  moreover then have "f (C \<triangle> f C) \<le> C \<triangle> f C"
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   491
    by (metis narrow1_acom narrow2_acom 1 mono order_trans)
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   492
  moreover have "n (C \<triangle> f C) < n C" using 1 2 by(simp add: n_narrow P_f)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   493
  ultimately show "(P (C \<triangle> f C) \<and> f (C \<triangle> f C) \<le> C \<triangle> f C) \<and> n(C \<triangle> f C) < n C"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   494
    by blast
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   495
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   496
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   497
locale Abs_Int_wn_measure = Abs_Int_wn where \<gamma>=\<gamma> + Measure_wn where m=m
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   498
  for \<gamma> :: "'av::{wn,bounded_lattice} \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"
49547
78be750222cf tuned termination proof
nipkow
parents: 49496
diff changeset
   499
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   500
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   501
subsubsection "Termination: Intervals"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   502
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   503
definition m_rep :: "eint2 \<Rightarrow> nat" where
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   504
"m_rep p = (if is_empty_rep p then 3 else
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   505
  let (l,h) = p in (case l of Minf \<Rightarrow> 0 | _ \<Rightarrow> 1) + (case h of Pinf \<Rightarrow> 0 | _ \<Rightarrow> 1))"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   506
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   507
lift_definition m_ivl :: "ivl \<Rightarrow> nat" is m_rep
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   508
by(auto simp: m_rep_def eq_ivl_iff)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   509
51924
e398ab28eaa7 standard ivl notation [l,h]
nipkow
parents: 51914
diff changeset
   510
lemma m_ivl_nice: "m_ivl[l,h] = (if [l,h] = \<bottom> then 3 else
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   511
   (if l = Minf then 0 else 1) + (if h = Pinf then 0 else 1))"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   512
unfolding bot_ivl_def
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   513
by transfer (auto simp: m_rep_def eq_ivl_empty split: extended.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   514
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   515
lemma m_ivl_height: "m_ivl iv \<le> 3"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   516
by transfer (simp add: m_rep_def split: prod.split extended.split)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   517
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   518
lemma m_ivl_anti_mono: "y \<le> x \<Longrightarrow> m_ivl x \<le> m_ivl y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   519
by transfer
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   520
   (auto simp: m_rep_def is_empty_rep_def \<gamma>_rep_cases le_iff_subset
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   521
         split: prod.split extended.splits if_splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   522
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   523
lemma m_ivl_widen:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   524
  "~ y \<le> x \<Longrightarrow> m_ivl(x \<nabla> y) < m_ivl x"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   525
by transfer
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   526
   (auto simp: m_rep_def widen_rep_def is_empty_rep_def \<gamma>_rep_cases le_iff_subset
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   527
         split: prod.split extended.splits if_splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   528
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   529
definition n_ivl :: "ivl \<Rightarrow> nat" where
51953
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   530
"n_ivl iv = 3 - m_ivl iv"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   531
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   532
lemma n_ivl_narrow:
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   533
  "x \<triangle> y < x \<Longrightarrow> n_ivl(x \<triangle> y) < n_ivl x"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   534
unfolding n_ivl_def
51385
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   535
apply(subst (asm) less_le_not_le)
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   536
apply transfer
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   537
by(auto simp add: m_rep_def narrow_rep_def is_empty_rep_def empty_rep_def \<gamma>_rep_cases le_iff_subset
f193d44d4918 termination proof for narrowing: fewer assumptions
nipkow
parents: 51372
diff changeset
   538
         split: prod.splits if_splits extended.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   539
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   540
61890
f6ded81f5690 abandoned attempt to unify sublocale and interpretation into global theories
haftmann
parents: 61671
diff changeset
   541
global_interpretation Abs_Int_wn_measure
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67019
diff changeset
   542
where \<gamma> = \<gamma>_ivl and num' = num_ivl and plus' = "(+)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   543
and test_num' = in_ivl
51974
9c80e62161a5 tuned names
nipkow
parents: 51953
diff changeset
   544
and inv_plus' = inv_plus_ivl and inv_less' = inv_less_ivl
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   545
and m = m_ivl and n = n_ivl and h = 3
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   546
proof (standard, goal_cases)
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   547
  case 2 thus ?case by(rule m_ivl_anti_mono)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   548
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   549
  case 1 thus ?case by(rule m_ivl_height)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   550
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   551
  case 3 thus ?case by(rule m_ivl_widen)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   552
next
61179
16775cad1a5c goali -> i
nipkow
parents: 55600
diff changeset
   553
  case 4 from 4(2) show ?case by(rule n_ivl_narrow)
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
   554
  \<comment> \<open>note that the first assms is unnecessary for intervals\<close>
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   555
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   556
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   557
lemma iter_winden_step_ivl_termination:
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   558
  "\<exists>C. iter_widen (step_ivl \<top>) (bot c) = Some C"
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   559
apply(rule iter_widen_termination[where m = "m_c" and P = "%C. strip C = c \<and> top_on_acom C (- vars C)"])
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   560
apply (auto simp add: m_c_widen top_on_bot top_on_step'[simplified comp_def vars_acom_def]
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   561
  vars_acom_def top_on_acom_widen)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   562
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   563
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   564
lemma iter_narrow_step_ivl_termination:
51953
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   565
  "top_on_acom C (- vars C) \<Longrightarrow> step_ivl \<top> C \<le> C \<Longrightarrow>
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   566
  \<exists>C'. iter_narrow (step_ivl \<top>) C = Some C'"
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   567
apply(rule iter_narrow_termination[where n = "n_c" and P = "%C'. strip C = strip C' \<and> top_on_acom C' (-vars C')"])
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   568
apply(auto simp: top_on_step'[simplified comp_def vars_acom_def]
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   569
        mono_step'_top n_c_narrow vars_acom_def top_on_acom_narrow)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   570
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   571
51953
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   572
theorem AI_wn_ivl_termination:
ae755fd6c883 tuned names
nipkow
parents: 51924
diff changeset
   573
  "\<exists>C. AI_wn_ivl c = Some C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   574
apply(auto simp: AI_wn_def pfp_wn_def iter_winden_step_ivl_termination
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   575
           split: option.split)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   576
apply(rule iter_narrow_step_ivl_termination)
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   577
apply(rule conjunct2)
51785
9685a5b1f7ce more standard order of arguments
nipkow
parents: 51722
diff changeset
   578
apply(rule iter_widen_inv[where f = "step' \<top>" and P = "%C. c = strip C & top_on_acom C (- vars C)"])
51711
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   579
apply(auto simp: top_on_acom_widen top_on_step'[simplified comp_def vars_acom_def]
df3426139651 complete revision: finally got rid of annoying L-predicate
nipkow
parents: 51390
diff changeset
   580
  iter_widen_pfp top_on_bot vars_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   581
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   582
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51385
diff changeset
   583
(*unused_thms Abs_Int_init - *)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   584
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   585
subsubsection "Counterexamples"
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   586
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   587
text\<open>Widening is increasing by assumption, but \<^prop>\<open>x \<le> f x\<close> is not an invariant of widening.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   588
It can already be lost after the first step:\<close>
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   589
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   590
lemma assumes "!!x y::'a::wn. x \<le> y \<Longrightarrow> f x \<le> f y"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   591
and "x \<le> f x" and "\<not> f x \<le> x" shows "x \<nabla> f x \<le> f(x \<nabla> f x)"
55357
1dd39517e1ce fight spurious nitpick timeouts
nipkow
parents: 53015
diff changeset
   592
nitpick[card = 3, expect = genuine, show_consts, timeout = 120]
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   593
(*
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   594
1 < 2 < 3,
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   595
f x = 2,
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   596
x widen y = 3 -- guarantees termination with top=3
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   597
x = 1
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   598
Now f is mono, x <= f x, not f x <= x
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   599
but x widen f x = 3, f 3 = 2, but not 3 <= 2
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   600
*)
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   601
oops
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   602
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   603
text\<open>Widening terminates but may converge more slowly than Kleene iteration.
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   604
In the following model, Kleene iteration goes from 0 to the least pfp
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67399
diff changeset
   605
in one step but widening takes 2 steps to reach a strictly larger pfp:\<close>
52504
52cd8bebc3b6 tuned names
nipkow
parents: 52019
diff changeset
   606
lemma assumes "!!x y::'a::wn. x \<le> y \<Longrightarrow> f x \<le> f y"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   607
and "x \<le> f x" and "\<not> f x \<le> x" and "f(f x) \<le> f x"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51245
diff changeset
   608
shows "f(x \<nabla> f x) \<le> x \<nabla> f x"
55357
1dd39517e1ce fight spurious nitpick timeouts
nipkow
parents: 53015
diff changeset
   609
nitpick[card = 4, expect = genuine, show_consts, timeout = 120]
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   610
(*
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   611
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   612
   0 < 1 < 2 < 3
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   613
f: 1   1   3   3
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   614
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   615
0 widen 1 = 2
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   616
2 widen 3 = 3
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   617
and x widen y arbitrary, eg 3, which guarantees termination
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   618
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   619
Kleene: f(f 0) = f 1 = 1 <= 1 = f 1
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   620
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   621
but
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   622
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   623
because not f 0 <= 0, we obtain 0 widen f 0 = 0 wide 1 = 2,
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   624
which is again not a pfp: not f 2 = 3 <= 2
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   625
Another widening step yields 2 widen f 2 = 2 widen 3 = 3
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   626
*)
49892
09956f7a00af proper 'oops' to force sequential checking here, and avoid spurious *** Interrupt stemming from crash of forked outer syntax element;
wenzelm
parents: 49579
diff changeset
   627
oops
49578
10f9f8608b4d added counterexamples
nipkow
parents: 49576
diff changeset
   628
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   629
end