src/HOL/IMP/Abs_Int1.thy
author nipkow
Mon, 11 Mar 2013 12:27:31 +0100
changeset 51390 1dff81cf425b
parent 51389 8a9f0503b1c0
child 51711 df3426139651
permissions -rw-r--r--
more factorisation of Step & Co
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     2
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     3
theory Abs_Int1
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     4
imports Abs_State
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     5
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
     6
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
     7
lemma le_iff_le_annos_zip: "C1 \<le> C2 \<longleftrightarrow>
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
     8
 (\<forall> (a1,a2) \<in> set(zip (annos C1) (annos C2)). a1 \<le> a2) \<and> strip C1 = strip C2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
     9
by(induct C1 C2 rule: less_eq_acom.induct) (auto simp: size_annos_same2)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    10
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    11
lemma le_iff_le_annos: "C1 \<le> C2 \<longleftrightarrow>
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    12
  strip C1 = strip C2 \<and> (\<forall> i<size(annos C1). annos C1 ! i \<le> annos C2 ! i)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    13
by(auto simp add: le_iff_le_annos_zip set_zip) (metis size_annos_same2)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    14
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    15
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    16
lemma mono_fun_L[simp]: "F \<in> L X \<Longrightarrow> F \<le> G \<Longrightarrow> x : X \<Longrightarrow> fun F x \<le> fun G x"
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    17
by(simp add: mono_fun L_st_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    18
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    19
lemma bot_in_L[simp]: "bot c \<in> L(vars c)"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    20
by(simp add: L_acom_def bot_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    21
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    22
lemma L_acom_simps[simp]: "SKIP {P} \<in> L X \<longleftrightarrow> P \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    23
  "(x ::= e {P}) \<in> L X \<longleftrightarrow> x : X \<and> vars e \<subseteq> X \<and> P \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    24
  "(C1;C2) \<in> L X \<longleftrightarrow> C1 \<in> L X \<and> C2 \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    25
  "(IF b THEN {P1} C1 ELSE {P2} C2 {Q}) \<in> L X \<longleftrightarrow>
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    26
   vars b \<subseteq> X \<and> C1 \<in> L X \<and> C2 \<in> L X \<and> P1 \<in> L X \<and> P2 \<in> L X \<and> Q \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    27
  "({I} WHILE b DO {P} C {Q}) \<in> L X \<longleftrightarrow>
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    28
   I \<in> L X \<and> vars b \<subseteq> X \<and> P \<in> L X \<and> C \<in> L X \<and> Q \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    29
by(auto simp add: L_acom_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    30
49353
nipkow
parents: 49344
diff changeset
    31
lemma post_in_annos: "post C : set(annos C)"
nipkow
parents: 49344
diff changeset
    32
by(induction C) auto
nipkow
parents: 49344
diff changeset
    33
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    34
lemma post_in_L[simp]: "C \<in> L X \<Longrightarrow> post C \<in> L X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    35
by(simp add: L_acom_def post_in_annos)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    36
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    37
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    38
subsection "Computable Abstract Interpretation"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    39
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    40
text{* Abstract interpretation over type @{text st} instead of functions. *}
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    41
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    42
context Gamma
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    43
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    44
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    45
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where
50896
nipkow
parents: 49547
diff changeset
    46
"aval' (N i) S = num' i" |
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    47
"aval' (V x) S = fun S x" |
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    48
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    49
49497
860b7c6bd913 tuned names
nipkow
parents: 49464
diff changeset
    50
lemma aval'_sound: "s : \<gamma>\<^isub>s S \<Longrightarrow> vars a \<subseteq> dom S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    51
by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    52
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    53
lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    54
  assumes "!!x e S. x : X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> S \<in> L X \<Longrightarrow> f1 x e (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (f2 x e S)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    55
          "!!b S. S \<in> L X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow> g1 b (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (g2 b S)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    56
  shows "C \<in> L X \<Longrightarrow> S \<in> L X \<Longrightarrow> Step f1 g1 (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (Step f2 g2 S C)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    57
proof(induction C arbitrary: S)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    58
qed (auto simp: assms intro!: mono_gamma_o post_in_L sup_ge1 sup_ge2)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    59
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    60
lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(update S x a)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    61
by(simp add: \<gamma>_st_def)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    62
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    63
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    64
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    65
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    66
lemma Step_in_L: fixes C :: "'a::semilatticeL acom"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    67
assumes "!!x e S. \<lbrakk>S \<in> L X; x \<in> X; vars e \<subseteq> X\<rbrakk> \<Longrightarrow> f x e S \<in> L X"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    68
        "!!b S. S \<in> L X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow> g b S \<in> L X"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    69
shows"\<lbrakk> C \<in> L X; S \<in> L X \<rbrakk> \<Longrightarrow> Step f g S C \<in> L X"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    70
proof(induction C arbitrary: S)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    71
  case Assign thus ?case
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    72
    by(auto simp: L_st_def assms split: option.splits)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    73
qed (auto simp: assms)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    74
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    75
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    76
text{* The for-clause (here and elsewhere) only serves the purpose of fixing
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    77
the name of the type parameter @{typ 'av} which would otherwise be renamed to
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    78
@{typ 'a}. *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    79
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
    80
locale Abs_Int = Gamma where \<gamma>=\<gamma> for \<gamma> :: "'av::semilattice \<Rightarrow> val set"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    81
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    82
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    83
definition "step' = Step
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    84
  (\<lambda>x e S. case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(update S x (aval' e S)))
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    85
  (\<lambda>b S. S)"
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
    86
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    87
definition AI :: "com \<Rightarrow> 'av st option acom option" where
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
    88
"AI c = pfp (step' (Top(vars c))) (bot c)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    89
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    90
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    91
lemma strip_step'[simp]: "strip(step' S C) = strip C"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    92
by(simp add: step'_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    93
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    94
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    95
text{* Soundness: *}
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
    96
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    97
lemma step_step': "C \<in> L X \<Longrightarrow> S \<in> L X \<Longrightarrow> step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    98
unfolding step_def step'_def
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
    99
by(rule gamma_Step_subcomm)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   100
  (auto simp: L_st_def intro!: aval'_sound in_gamma_update split: option.splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   101
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   102
lemma step'_in_L[simp]: "\<lbrakk> C \<in> L X; S \<in> L X \<rbrakk> \<Longrightarrow> step' S C \<in> L X"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   103
unfolding step'_def
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   104
by(rule Step_in_L)(auto simp: L_st_def step'_def split: option.splits)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   105
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
   106
lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   107
proof(simp add: CS_def AI_def)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   108
  assume 1: "pfp (step' (Top(vars c))) (bot c) = Some C"
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   109
  have "C \<in> L(vars c)"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   110
    by(rule pfp_inv[where P = "%C. C \<in> L(vars c)", OF 1 _ bot_in_L])
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   111
      (erule step'_in_L[OF _ Top_in_L])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   112
  have pfp': "step' (Top(vars c)) C \<le> C" by(rule pfp_pfp[OF 1])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   113
  have 2: "step (\<gamma>\<^isub>o(Top(vars c))) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
   114
  proof(rule order_trans)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   115
    show "step (\<gamma>\<^isub>o (Top(vars c))) (\<gamma>\<^isub>c C) \<le>  \<gamma>\<^isub>c (step' (Top(vars c)) C)"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   116
      by(rule step_step'[OF `C \<in> L(vars c)` Top_in_L])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   117
    show "\<gamma>\<^isub>c (step' (Top(vars c)) C) \<le> \<gamma>\<^isub>c C"
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
   118
      by(rule mono_gamma_c[OF pfp'])
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   119
  qed
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   120
  have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   121
  have "lfp c (step (\<gamma>\<^isub>o(Top(vars c)))) \<le> \<gamma>\<^isub>c C"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   122
    by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o(Top(vars c)))", OF 3 2])
50986
c54ea7f5418f simplified proofs
nipkow
parents: 50896
diff changeset
   123
  thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   124
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   125
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   126
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   127
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   128
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   129
subsubsection "Monotonicity"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   130
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
   131
lemma le_sup_disj: "y \<in> L X \<Longrightarrow> (z::_::semilatticeL) \<in> L X \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   132
  x \<le> y \<or> x \<le> z \<Longrightarrow> x \<le> y \<squnion> z"
51389
8a9f0503b1c0 factored out Step
nipkow
parents: 51372
diff changeset
   133
by (metis sup_ge1 sup_ge2 order_trans)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   134
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   135
theorem mono2_Step: fixes C1 :: "'a::semilatticeL acom"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   136
  assumes "!!x e S1 S2. \<lbrakk>S1 \<in> L X; S2 \<in> L X; x \<in> X; vars e \<subseteq> X; S1 \<le> S2\<rbrakk> \<Longrightarrow> f x e S1 \<le> f x e S2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   137
          "!!b S1 S2. S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> vars b \<subseteq> X \<Longrightarrow> S1 \<le> S2 \<Longrightarrow> g b S1 \<le> g b S2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   138
  shows "S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> C1 \<in> L X \<Longrightarrow> C2 \<in> L X \<Longrightarrow>
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   139
  S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> Step f g S1 C1 \<le> Step f g S2 C2"
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   140
by(induction C1 C2 arbitrary: S1 S2 rule: less_eq_acom.induct)
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   141
  (auto simp: mono_post assms le_sup_disj le_sup_disj[OF  post_in_L post_in_L])
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   142
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   143
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   144
locale Abs_Int_mono = Abs_Int +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   145
assumes mono_plus': "a1 \<le> b1 \<Longrightarrow> a2 \<le> b2 \<Longrightarrow> plus' a1 a2 \<le> plus' b1 b2"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   146
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   147
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   148
lemma mono_aval':
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   149
  "S1 \<le> S2 \<Longrightarrow> S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> vars e \<subseteq> X \<Longrightarrow> aval' e S1 \<le> aval' e S2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   150
by(induction e) (auto simp: less_eq_st_def mono_plus' L_st_def)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   151
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   152
theorem mono_step': "S1 \<in> L X \<Longrightarrow> S2 \<in> L X \<Longrightarrow> C1 \<in> L X \<Longrightarrow> C2 \<in> L X \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   153
  S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> step' S1 C1 \<le> step' S2 C2"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   154
unfolding step'_def
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   155
by(rule mono2_Step) (auto simp: mono_aval' split: option.split)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   156
51036
e7b54119c436 tuned top
nipkow
parents: 50986
diff changeset
   157
lemma mono_step'_top: "C \<in> L X \<Longrightarrow> C' \<in> L X \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   158
  C \<le> C' \<Longrightarrow> step' (Top X) C \<le> step' (Top X) C'"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   159
by (metis Top_in_L mono_step' order_refl)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   160
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   161
lemma pfp_bot_least:
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   162
assumes "\<forall>x\<in>L(vars c)\<inter>{C. strip C = c}.\<forall>y\<in>L(vars c)\<inter>{C. strip C = c}.
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   163
  x \<le> y \<longrightarrow> f x \<le> f y"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   164
and "\<forall>C. C \<in> L(vars c)\<inter>{C. strip C = c} \<longrightarrow> f C \<in> L(vars c)\<inter>{C. strip C = c}"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   165
and "f C' \<le> C'" "strip C' = c" "C' \<in> L(vars c)" "pfp f (bot c) = Some C"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   166
shows "C \<le> C'"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   167
by(rule while_least[OF assms(1,2) _ _ assms(3) _ assms(6)[unfolded pfp_def]])
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   168
  (simp_all add: assms(4,5) bot_least)
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   169
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   170
lemma AI_least_pfp: assumes "AI c = Some C"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   171
and "step' (Top(vars c)) C' \<le> C'" "strip C' = c" "C' \<in> L(vars c)"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   172
shows "C \<le> C'"
51390
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   173
by(rule pfp_bot_least[OF _ _ assms(2-4) assms(1)[unfolded AI_def]])
1dff81cf425b more factorisation of Step & Co
nipkow
parents: 51389
diff changeset
   174
  (simp_all add: mono_step'_top)
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   175
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   176
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   177
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   178
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   179
subsubsection "Termination"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   180
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   181
lemma pfp_termination:
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   182
fixes x0 :: "'a::order" and m :: "'a \<Rightarrow> nat"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   183
assumes mono: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   184
and m: "\<And>x y. I x \<Longrightarrow> I y \<Longrightarrow> x < y \<Longrightarrow> m x > m y"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   185
and I: "\<And>x y. I x \<Longrightarrow> I(f x)" and "I x0" and "x0 \<le> f x0"
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   186
shows "\<exists>x. pfp f x0 = Some x"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   187
proof(simp add: pfp_def, rule wf_while_option_Some[where P = "%x. I x & x \<le> f x"])
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   188
  show "wf {(y,x). ((I x \<and> x \<le> f x) \<and> \<not> f x \<le> x) \<and> y = f x}"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   189
    by(rule wf_subset[OF wf_measure[of m]]) (auto simp: m I)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   190
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   191
  show "I x0 \<and> x0 \<le> f x0" using `I x0` `x0 \<le> f x0` by blast
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   192
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   193
  fix x assume "I x \<and> x \<le> f x" thus "I(f x) \<and> f x \<le> f(f x)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   194
    by (blast intro: I mono)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   195
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   196
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   197
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   198
locale Measure1 =
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   199
fixes m :: "'av::order \<Rightarrow> nat"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   200
fixes h :: "nat"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   201
assumes h: "m x \<le> h"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   202
begin
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   203
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   204
definition m_s :: "'av st \<Rightarrow> nat" ("m\<^isub>s") where
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   205
"m_s S = (\<Sum> x \<in> dom S. m(fun S x))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   206
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   207
lemma m_s_h: "x \<in> L X \<Longrightarrow> finite X \<Longrightarrow> m_s x \<le> h * card X"
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   208
by(simp add: L_st_def m_s_def)
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   209
  (metis nat_mult_commute of_nat_id setsum_bounded[OF h])
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   210
49432
3f4104ccca77 beautified names
nipkow
parents: 49431
diff changeset
   211
definition m_o :: "nat \<Rightarrow> 'av st option \<Rightarrow> nat" ("m\<^isub>o") where
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   212
"m_o d opt = (case opt of None \<Rightarrow> h*d+1 | Some S \<Rightarrow> m_s S)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   213
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   214
lemma m_o_h: "ost \<in> L X \<Longrightarrow> finite X \<Longrightarrow> m_o (card X) ost \<le> (h*card X + 1)"
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   215
by(auto simp add: m_o_def m_s_h split: option.split dest!:m_s_h)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   216
49432
3f4104ccca77 beautified names
nipkow
parents: 49431
diff changeset
   217
definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^isub>c") where
49431
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   218
"m_c C = (\<Sum>i<size(annos C). m_o (card(vars(strip C))) (annos C ! i))"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   219
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   220
lemma m_c_h: assumes "C \<in> L(vars(strip C))"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   221
shows "m_c C \<le> size(annos C) * (h * card(vars(strip C)) + 1)"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   222
proof-
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   223
  let ?X = "vars(strip C)" let ?n = "card ?X" let ?a = "size(annos C)"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   224
  { fix i assume "i < ?a"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   225
    hence "annos C ! i \<in> L ?X" using assms by(simp add: L_acom_def)
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   226
    note m_o_h[OF this finite_cvars]
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   227
  } note 1 = this
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   228
  have "m_c C = (\<Sum>i<?a. m_o ?n (annos C ! i))" by(simp add: m_c_def)
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   229
  also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   230
    apply(rule setsum_mono) using 1 by simp
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   231
  also have "\<dots> = ?a * (h * ?n + 1)" by simp
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   232
  finally show ?thesis .
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   233
qed
bf491a1c15c2 proved all upper bounds
nipkow
parents: 49402
diff changeset
   234
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   235
end
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   236
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   237
locale Measure = Measure1 +
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   238
assumes m2: "x < y \<Longrightarrow> m x > m y"
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   239
begin
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   240
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   241
lemma m1: "x \<le> y \<Longrightarrow> m x \<ge> m y"
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   242
by(auto simp: le_less m2)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   243
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   244
lemma m_s2: "finite(dom S1) \<Longrightarrow> S1 < S2 \<Longrightarrow> m_s S1 > m_s S2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   245
proof(auto simp add: less_st_def less_eq_st_def m_s_def)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   246
  assume "finite(dom S2)" and 0: "\<forall>x\<in>dom S2. fun S1 x \<le> fun S2 x"
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   247
  hence 1: "\<forall>x\<in>dom S2. m(fun S1 x) \<ge> m(fun S2 x)" by (metis m1)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   248
  fix x assume "x \<in> dom S2" "\<not> fun S2 x \<le> fun S1 x"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   249
  hence 2: "\<exists>x\<in>dom S2. m(fun S1 x) > m(fun S2 x)" by (metis 0 m2 less_le_not_le)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   250
  from setsum_strict_mono_ex1[OF `finite(dom S2)` 1 2]
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   251
  show "(\<Sum>x\<in>dom S2. m (fun S2 x)) < (\<Sum>x\<in>dom S2. m (fun S1 x))" .
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   252
qed
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   253
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   254
lemma m_o2: "finite X \<Longrightarrow> o1 \<in> L X \<Longrightarrow> o2 \<in> L X \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   255
  o1 < o2 \<Longrightarrow> m_o (card X) o1 > m_o (card X) o2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   256
proof(induction o1 o2 rule: less_eq_option.induct)
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   257
  case 1 thus ?case by (auto simp: m_o_def L_st_def m_s2 less_option_def)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   258
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   259
  case 2 thus ?case by(auto simp: m_o_def less_option_def le_imp_less_Suc m_s_h)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   260
next
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   261
  case 3 thus ?case by (auto simp: less_option_def)
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   262
qed
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   263
51372
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   264
lemma m_o1: "finite X \<Longrightarrow> o1 \<in> L X \<Longrightarrow> o2 \<in> L X \<Longrightarrow>
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   265
  o1 \<le> o2 \<Longrightarrow> m_o (card X) o1 \<ge> m_o (card X) o2"
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   266
by(auto simp: le_less m_o2)
d315e9a9ee72 simplified basic termination proof
nipkow
parents: 51359
diff changeset
   267
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   268
lemma m_c2: "C1 \<in> L(vars(strip C1)) \<Longrightarrow> C2 \<in> L(vars(strip C2)) \<Longrightarrow>
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   269
  C1 < C2 \<Longrightarrow> m_c C1 > m_c C2"
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   270
proof(auto simp add: le_iff_le_annos m_c_def size_annos_same[of C1 C2] less_acom_def L_acom_def)
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   271
  let ?X = "vars(strip C2)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   272
  let ?n = "card ?X"
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   273
  assume V1: "\<forall>a\<in>set(annos C1). a \<in> L ?X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   274
    and V2: "\<forall>a\<in>set(annos C2). a \<in> L ?X"
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   275
    and strip_eq: "strip C1 = strip C2"
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   276
    and 0: "\<forall>i<size(annos C2). annos C1 ! i \<le> annos C2 ! i"
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   277
  hence 1: "\<forall>i<size(annos C2). m_o ?n (annos C1 ! i) \<ge> m_o ?n (annos C2 ! i)"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   278
    by (auto simp: all_set_conv_all_nth)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   279
       (metis finite_cvars m_o1 size_annos_same2)
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   280
  fix i assume "i < size(annos C2)" "\<not> annos C2 ! i \<le> annos C1 ! i"
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   281
  hence "m_o ?n (annos C1 ! i) > m_o ?n (annos C2 ! i)" (is "?P i")
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   282
    by(metis m_o2[OF finite_cvars] V1 V2 nth_mem size_annos_same[OF strip_eq] 0 less_option_def)
49396
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   283
  hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   284
  show "(\<Sum>i<size(annos C2). m_o ?n (annos C2 ! i))
73fb17ed2e08 converted wt into a set, tuned names
nipkow
parents: 49353
diff changeset
   285
         < (\<Sum>i<size(annos C2). m_o ?n (annos C1 ! i))"
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   286
    apply(rule setsum_strict_mono_ex1) using 1 2 by (auto)
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   287
qed
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   288
49547
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   289
end
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   290
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   291
locale Abs_Int_measure =
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   292
  Abs_Int_mono where \<gamma>=\<gamma> + Measure where m=m
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   293
  for \<gamma> :: "'av::semilattice \<Rightarrow> val set" and m :: "'av \<Rightarrow> nat"
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   294
begin
78be750222cf tuned termination proof
nipkow
parents: 49497
diff changeset
   295
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   296
lemma AI_Some_measure: "\<exists>C. AI c = Some C"
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   297
unfolding AI_def
51359
00b45c7e831f major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents: 51036
diff changeset
   298
apply(rule pfp_termination[where I = "\<lambda>C. strip C = c \<and> C \<in> L(vars c)" and m="m_c"])
49464
4e4bdd12280f removed lpfp and proved least pfp thm
nipkow
parents: 49432
diff changeset
   299
apply(simp_all add: m_c2 mono_step'_top bot_least)
47613
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   300
done
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   301
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   302
end
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   303
e72e44cee6f2 added revised version of Abs_Int
nipkow
parents:
diff changeset
   304
end