src/HOL/IMP/Live_True.thy
author wenzelm
Fri, 22 Jun 2018 20:31:49 +0200
changeset 68484 59793df7f853
parent 67406 23307fd33906
child 68776 403dd13cf6e9
permissions -rw-r--r--
clarified document antiquotation @{theory};
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     2
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     3
theory Live_True
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 53015
diff changeset
     4
imports "HOL-Library.While_Combinator" Vars Big_Step
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     5
begin
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     6
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     7
subsection "True Liveness Analysis"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     8
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
     9
fun L :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    10
"L SKIP X = X" |
51436
790310525e97 tuned (in particular bold fonts)
nipkow
parents: 50180
diff changeset
    11
"L (x ::= a) X = (if x \<in> X then vars a \<union> (X - {x}) else X)" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52072
diff changeset
    12
"L (c\<^sub>1;; c\<^sub>2) X = L c\<^sub>1 (L c\<^sub>2 X)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52072
diff changeset
    13
"L (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = vars b \<union> L c\<^sub>1 X \<union> L c\<^sub>2 X" |
50009
nipkow
parents: 50007
diff changeset
    14
"L (WHILE b DO c) X = lfp(\<lambda>Y. vars b \<union> X \<union> L c Y)"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    15
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    16
lemma L_mono: "mono (L c)"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    17
proof-
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    18
  have "X \<subseteq> Y \<Longrightarrow> L c X \<subseteq> L c Y" for X Y
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    19
  proof(induction c arbitrary: X Y)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    20
    case (While b c)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    21
    show ?case
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    22
    proof(simp, rule lfp_mono)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    23
      fix Z show "vars b \<union> X \<union> L c Z \<subseteq> vars b \<union> Y \<union> L c Z"
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    24
        using While by auto
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    25
    qed
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    26
  next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    27
    case If thus ?case by(auto simp: subset_iff)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    28
  qed auto
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 66453
diff changeset
    29
  thus ?thesis by(rule monoI)
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    30
qed
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    31
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    32
lemma mono_union_L:
50009
nipkow
parents: 50007
diff changeset
    33
  "mono (\<lambda>Y. X \<union> L c Y)"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    34
by (metis (no_types) L_mono mono_def order_eq_iff set_eq_subset sup_mono)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    35
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    36
lemma L_While_unfold:
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    37
  "L (WHILE b DO c) X = vars b \<union> X \<union> L c (L (WHILE b DO c) X)"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    38
by(metis lfp_unfold[OF mono_union_L] L.simps(5))
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    39
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    40
lemma L_While_pfp: "L c (L (WHILE b DO c) X) \<subseteq> L (WHILE b DO c) X"
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    41
using L_While_unfold by blast
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    42
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    43
lemma L_While_vars: "vars b \<subseteq> L (WHILE b DO c) X"
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    44
using L_While_unfold by blast
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    45
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    46
lemma L_While_X: "X \<subseteq> L (WHILE b DO c) X"
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    47
using L_While_unfold by blast
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    48
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    49
text\<open>Disable @{text "L WHILE"} equation and reason only with @{text "L WHILE"} constraints:\<close>
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    50
declare L.simps(5)[simp del]
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    51
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    52
51975
7bab3fb52e3e tuned names
nipkow
parents: 51698
diff changeset
    53
subsection "Correctness"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    54
51975
7bab3fb52e3e tuned names
nipkow
parents: 51698
diff changeset
    55
theorem L_correct:
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    56
  "(c,s) \<Rightarrow> s'  \<Longrightarrow> s = t on L c X \<Longrightarrow>
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    57
  \<exists> t'. (c,t) \<Rightarrow> t' & s' = t' on X"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    58
proof (induction arbitrary: X t rule: big_step_induct)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    59
  case Skip then show ?case by auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    60
next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    61
  case Assign then show ?case
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    62
    by (auto simp: ball_Un)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    63
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 46365
diff changeset
    64
  case (Seq c1 s1 s2 c2 s3 X t1)
151d137f1095 renamed Semi to Seq
nipkow
parents: 46365
diff changeset
    65
  from Seq.IH(1) Seq.prems obtain t2 where
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    66
    t12: "(c1, t1) \<Rightarrow> t2" and s2t2: "s2 = t2 on L c2 X"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    67
    by simp blast
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 46365
diff changeset
    68
  from Seq.IH(2)[OF s2t2] obtain t3 where
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    69
    t23: "(c2, t2) \<Rightarrow> t3" and s3t3: "s3 = t3 on X"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    70
    by auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    71
  show ?case using t12 t23 s3t3 by auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    72
next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    73
  case (IfTrue b s c1 s' c2)
50009
nipkow
parents: 50007
diff changeset
    74
  hence "s = t on vars b" and "s = t on L c1 X" by auto
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    75
  from  bval_eq_if_eq_on_vars[OF this(1)] IfTrue(1) have "bval b t" by simp
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    76
  from IfTrue.IH[OF \<open>s = t on L c1 X\<close>] obtain t' where
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    77
    "(c1, t) \<Rightarrow> t'" "s' = t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    78
  thus ?case using \<open>bval b t\<close> by auto
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    79
next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    80
  case (IfFalse b s c2 s' c1)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    81
  hence "s = t on vars b" "s = t on L c2 X" by auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    82
  from  bval_eq_if_eq_on_vars[OF this(1)] IfFalse(1) have "~bval b t" by simp
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    83
  from IfFalse.IH[OF \<open>s = t on L c2 X\<close>] obtain t' where
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    84
    "(c2, t) \<Rightarrow> t'" "s' = t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    85
  thus ?case using \<open>~bval b t\<close> by auto
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    86
next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    87
  case (WhileFalse b s c)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    88
  hence "~ bval b t"
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    89
    by (metis L_While_vars bval_eq_if_eq_on_vars set_mp)
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    90
  thus ?case using WhileFalse.prems L_While_X[of X b c] by auto
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    91
next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    92
  case (WhileTrue b s1 c s2 s3 X t1)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    93
  let ?w = "WHILE b DO c"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    94
  from \<open>bval b s1\<close> WhileTrue.prems have "bval b t1"
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    95
    by (metis L_While_vars bval_eq_if_eq_on_vars set_mp)
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
    96
  have "s1 = t1 on L c (L ?w X)" using  L_While_pfp WhileTrue.prems
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    97
    by (blast)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    98
  from WhileTrue.IH(1)[OF this] obtain t2 where
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
    99
    "(c, t1) \<Rightarrow> t2" "s2 = t2 on L ?w X" by auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   100
  from WhileTrue.IH(2)[OF this(2)] obtain t3 where "(?w,t2) \<Rightarrow> t3" "s3 = t3 on X"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   101
    by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   102
  with \<open>bval b t1\<close> \<open>(c, t1) \<Rightarrow> t2\<close> show ?case by auto
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   103
qed
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   104
50009
nipkow
parents: 50007
diff changeset
   105
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   106
subsection "Executability"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   107
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   108
lemma L_subset_vars: "L c X \<subseteq> rvars c \<union> X"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   109
proof(induction c arbitrary: X)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   110
  case (While b c)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   111
  have "lfp(\<lambda>Y. vars b \<union> X \<union> L c Y) \<subseteq> vars b \<union> rvars c \<union> X"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   112
    using While.IH[of "vars b \<union> rvars c \<union> X"]
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   113
    by (auto intro!: lfp_lowerbound)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
   114
  thus ?case by (simp add: L.simps(5))
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   115
qed auto
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   116
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   117
text\<open>Make @{const L} executable by replacing @{const lfp} with the @{const
68484
59793df7f853 clarified document antiquotation @{theory};
wenzelm
parents: 67406
diff changeset
   118
while} combinator from theory @{theory "HOL-Library.While_Combinator"}. The @{const while}
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   119
combinator obeys the recursion equation
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   120
@{thm[display] While_Combinator.while_unfold[no_vars]}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   121
and is thus executable.\<close>
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   122
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   123
lemma L_While: fixes b c X
51464
nipkow
parents: 51460
diff changeset
   124
assumes "finite X" defines "f == \<lambda>Y. vars b \<union> X \<union> L c Y"
nipkow
parents: 51460
diff changeset
   125
shows "L (WHILE b DO c) X = while (\<lambda>Y. f Y \<noteq> Y) f {}" (is "_ = ?r")
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   126
proof -
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   127
  let ?V = "vars b \<union> rvars c \<union> X"
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   128
  have "lfp f = ?r"
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   129
  proof(rule lfp_while[where C = "?V"])
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   130
    show "mono f" by(simp add: f_def mono_union_L)
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   131
  next
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   132
    fix Y show "Y \<subseteq> ?V \<Longrightarrow> f Y \<subseteq> ?V"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   133
      unfolding f_def using L_subset_vars[of c] by blast
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   134
  next
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   135
    show "finite ?V" using \<open>finite X\<close> by simp
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   136
  qed
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
   137
  thus ?thesis by (simp add: f_def L.simps(5))
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   138
qed
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   139
51464
nipkow
parents: 51460
diff changeset
   140
lemma L_While_let: "finite X \<Longrightarrow> L (WHILE b DO c) X =
nipkow
parents: 51460
diff changeset
   141
  (let f = (\<lambda>Y. vars b \<union> X \<union> L c Y)
nipkow
parents: 51460
diff changeset
   142
   in while (\<lambda>Y. f Y \<noteq> Y) f {})"
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
   143
by(simp add: L_While)
51464
nipkow
parents: 51460
diff changeset
   144
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   145
lemma L_While_set: "L (WHILE b DO c) (set xs) =
51464
nipkow
parents: 51460
diff changeset
   146
  (let f = (\<lambda>Y. vars b \<union> set xs \<union> L c Y)
nipkow
parents: 51460
diff changeset
   147
   in while (\<lambda>Y. f Y \<noteq> Y) f {})"
nipkow
parents: 51460
diff changeset
   148
by(rule L_While_let, simp)
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   149
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   150
text\<open>Replace the equation for @{text "L (WHILE \<dots>)"} by the executable @{thm[source] L_While_set}:\<close>
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   151
lemmas [code] = L.simps(1-4) L_While_set
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   152
text\<open>Sorry, this syntax is odd.\<close>
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   153
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   154
text\<open>A test:\<close>
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51975
diff changeset
   155
lemma "(let b = Less (N 0) (V ''y''); c = ''y'' ::= V ''x'';; ''x'' ::= V ''z''
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   156
  in L (WHILE b DO c) {''y''}) = {''x'', ''y'', ''z''}"
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   157
by eval
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   158
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   159
50009
nipkow
parents: 50007
diff changeset
   160
subsection "Limiting the number of iterations"
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   161
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   162
text\<open>The final parameter is the default value:\<close>
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   163
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   164
fun iter :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   165
"iter f 0 p d = d" |
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   166
"iter f (Suc n) p d = (if f p = p then p else iter f n (f p) d)"
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   167
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   168
text\<open>A version of @{const L} with a bounded number of iterations (here: 2)
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   169
in the WHILE case:\<close>
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   170
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   171
fun Lb :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   172
"Lb SKIP X = X" |
50009
nipkow
parents: 50007
diff changeset
   173
"Lb (x ::= a) X = (if x \<in> X then X - {x} \<union> vars a else X)" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52072
diff changeset
   174
"Lb (c\<^sub>1;; c\<^sub>2) X = (Lb c\<^sub>1 \<circ> Lb c\<^sub>2) X" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52072
diff changeset
   175
"Lb (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = vars b \<union> Lb c\<^sub>1 X \<union> Lb c\<^sub>2 X" |
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   176
"Lb (WHILE b DO c) X = iter (\<lambda>A. vars b \<union> X \<union> Lb c A) 2 {} (vars b \<union> rvars c \<union> X)"
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   177
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
   178
text\<open>@{const Lb} (and @{const iter}) is not monotone!\<close>
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51975
diff changeset
   179
lemma "let w = WHILE Bc False DO (''x'' ::= V ''y'';; ''z'' ::= V ''x'')
50009
nipkow
parents: 50007
diff changeset
   180
  in \<not> (Lb w {''z''} \<subseteq> Lb w {''y'',''z''})"
nipkow
parents: 50007
diff changeset
   181
by eval
nipkow
parents: 50007
diff changeset
   182
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   183
lemma lfp_subset_iter:
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   184
  "\<lbrakk> mono f; !!X. f X \<subseteq> f' X; lfp f \<subseteq> D \<rbrakk> \<Longrightarrow> lfp f \<subseteq> iter f' n A D"
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   185
proof(induction n arbitrary: A)
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   186
  case 0 thus ?case by simp
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   187
next
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   188
  case Suc thus ?case by simp (metis lfp_lowerbound)
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   189
qed
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   190
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   191
lemma "L c X \<subseteq> Lb c X"
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   192
proof(induction c arbitrary: X)
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   193
  case (While b c)
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   194
  let ?f  = "\<lambda>A. vars b \<union> X \<union> L  c A"
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   195
  let ?fb = "\<lambda>A. vars b \<union> X \<union> Lb c A"
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   196
  show ?case
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51464
diff changeset
   197
  proof (simp add: L.simps(5), rule lfp_subset_iter[OF mono_union_L])
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   198
    show "!!X. ?f X \<subseteq> ?fb X" using While.IH by blast
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   199
    show "lfp ?f \<subseteq> vars b \<union> rvars c \<union> X"
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   200
      by (metis (full_types) L.simps(5) L_subset_vars rvars.simps(5))
50007
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   201
  qed
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   202
next
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   203
  case Seq thus ?case by simp (metis (full_types) L_mono monoD subset_trans)
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   204
qed auto
56f269baae76 executable true liveness analysis incl an approximating version
nipkow
parents: 48256
diff changeset
   205
45812
0b02adadf384 added IMP/Live_True.thy
nipkow
parents:
diff changeset
   206
end