src/HOL/IMP/Live.thy
author wenzelm
Fri, 08 Dec 2023 13:09:39 +0100
changeset 79198 2586c8b422ed
parent 69712 dc85b5b3a532
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     2
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 53015
diff changeset
     3
section "Live Variable Analysis"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     4
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     5
theory Live imports Vars Big_Step
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     6
begin
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     7
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     8
subsection "Liveness Analysis"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
     9
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45015
diff changeset
    10
fun L :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
51425
nipkow
parents: 51396
diff changeset
    11
"L SKIP X = X" |
51448
nipkow
parents: 51433
diff changeset
    12
"L (x ::= a) X = vars a \<union> (X - {x})" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
    13
"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: 52046
diff changeset
    14
"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" |
51425
nipkow
parents: 51396
diff changeset
    15
"L (WHILE b DO c) X = vars b \<union> X \<union> L c X"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    16
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51975
diff changeset
    17
value "show (L (''y'' ::= V ''z'';; ''x'' ::= Plus (V ''y'') (V ''z'')) {''x''})"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    18
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    19
value "show (L (WHILE Less (V ''x'') (V ''x'') DO ''y'' ::= V ''z'') {''x''})"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    20
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45015
diff changeset
    21
fun "kill" :: "com \<Rightarrow> vname set" where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    22
"kill SKIP = {}" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    23
"kill (x ::= a) = {x}" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
    24
"kill (c\<^sub>1;; c\<^sub>2) = kill c\<^sub>1 \<union> kill c\<^sub>2" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
    25
"kill (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = kill c\<^sub>1 \<inter> kill c\<^sub>2" |
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    26
"kill (WHILE b DO c) = {}"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    27
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45015
diff changeset
    28
fun gen :: "com \<Rightarrow> vname set" where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    29
"gen SKIP = {}" |
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    30
"gen (x ::= a) = vars a" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
    31
"gen (c\<^sub>1;; c\<^sub>2) = gen c\<^sub>1 \<union> (gen c\<^sub>2 - kill c\<^sub>1)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
    32
"gen (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = vars b \<union> gen c\<^sub>1 \<union> gen c\<^sub>2" |
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    33
"gen (WHILE b DO c) = vars b \<union> gen c"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    34
51433
nipkow
parents: 51425
diff changeset
    35
lemma L_gen_kill: "L c X = gen c \<union> (X - kill c)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    36
by(induct c arbitrary:X) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    37
45771
a70465244096 added lemmas
nipkow
parents: 45770
diff changeset
    38
lemma L_While_pfp: "L c (L (WHILE b DO c) X) \<subseteq> L (WHILE b DO c) X"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    39
by(auto simp add:L_gen_kill)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    40
45771
a70465244096 added lemmas
nipkow
parents: 45770
diff changeset
    41
lemma L_While_lpfp:
45784
nipkow
parents: 45771
diff changeset
    42
  "vars b \<union> X \<union> L c P \<subseteq> P \<Longrightarrow> L (WHILE b DO c) X \<subseteq> P"
45771
a70465244096 added lemmas
nipkow
parents: 45770
diff changeset
    43
by(simp add: L_gen_kill)
a70465244096 added lemmas
nipkow
parents: 45770
diff changeset
    44
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    45
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: 51448
diff changeset
    46
by auto
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    47
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    48
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: 51448
diff changeset
    49
by auto
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    50
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    51
text\<open>Disable L WHILE equation and reason only with L WHILE constraints\<close>
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    52
declare L.simps(5)[simp del]
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    53
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
    54
subsection "Correctness"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    55
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
    56
theorem L_correct:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    57
  "(c,s) \<Rightarrow> s'  \<Longrightarrow> s = t on L c X \<Longrightarrow>
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    58
  \<exists> t'. (c,t) \<Rightarrow> t' & s' = t' on X"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
    59
proof (induction arbitrary: X t rule: big_step_induct)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    60
  case Skip then show ?case by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    61
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    62
  case Assign then show ?case
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    63
    by (auto simp: ball_Un)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    64
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
    65
  case (Seq c1 s1 s2 c2 s3 X t1)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
    66
  from Seq.IH(1) Seq.prems obtain t2 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    67
    t12: "(c1, t1) \<Rightarrow> t2" and s2t2: "s2 = t2 on L c2 X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    68
    by simp blast
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
    69
  from Seq.IH(2)[OF s2t2] obtain t3 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    70
    t23: "(c2, t2) \<Rightarrow> t3" and s3t3: "s3 = t3 on X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    71
    by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    72
  show ?case using t12 t23 s3t3 by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    73
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    74
  case (IfTrue b s c1 s' c2)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    75
  hence "s = t on vars b" "s = t on L c1 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    76
  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: 58889
diff changeset
    77
  from IfTrue.IH[OF \<open>s = t on L c1 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    78
    "(c1, t) \<Rightarrow> t'" "s' = t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    79
  thus ?case using \<open>bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    80
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    81
  case (IfFalse b s c2 s' c1)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    82
  hence "s = t on vars b" "s = t on L c2 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    83
  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: 58889
diff changeset
    84
  from IfFalse.IH[OF \<open>s = t on L c2 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    85
    "(c2, t) \<Rightarrow> t'" "s' = t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    86
  thus ?case using \<open>~bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    87
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    88
  case (WhileFalse b s c)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    89
  hence "~ bval b t"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    90
    by (metis L_While_vars bval_eq_if_eq_on_vars subsetD)
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    91
  thus ?case by(metis WhileFalse.prems L_While_X big_step.WhileFalse subsetD)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    92
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    93
  case (WhileTrue b s1 c s2 s3 X t1)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    94
  let ?w = "WHILE b DO c"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
    95
  from \<open>bval b s1\<close> WhileTrue.prems have "bval b t1"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    96
    by (metis L_While_vars bval_eq_if_eq_on_vars subsetD)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
    97
  have "s1 = t1 on L c (L ?w X)" using L_While_pfp WhileTrue.prems
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
    98
    by (blast)
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
    99
  from WhileTrue.IH(1)[OF this] obtain t2 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   100
    "(c, t1) \<Rightarrow> t2" "s2 = t2 on L ?w X" by auto
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   101
  from WhileTrue.IH(2)[OF this(2)] obtain t3 where "(?w,t2) \<Rightarrow> t3" "s3 = t3 on X"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   102
    by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   103
  with \<open>bval b t1\<close> \<open>(c, t1) \<Rightarrow> t2\<close> show ?case by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   104
qed
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   105
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   106
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   107
subsection "Program Optimization"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   108
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   109
text\<open>Burying assignments to dead variables:\<close>
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45015
diff changeset
   110
fun bury :: "com \<Rightarrow> vname set \<Rightarrow> com" where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   111
"bury SKIP X = SKIP" |
50009
nipkow
parents: 47818
diff changeset
   112
"bury (x ::= a) X = (if x \<in> X then x ::= a else SKIP)" |
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
   113
"bury (c\<^sub>1;; c\<^sub>2) X = (bury c\<^sub>1 (L c\<^sub>2 X);; bury c\<^sub>2 X)" |
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
   114
"bury (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = IF b THEN bury c\<^sub>1 X ELSE bury c\<^sub>2 X" |
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   115
"bury (WHILE b DO c) X = WHILE b DO bury c (L (WHILE b DO c) X)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   116
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   117
text\<open>We could prove the analogous lemma to @{thm[source]L_correct}, and the
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   118
proof would be very similar. However, we phrase it as a semantics
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   119
preservation property:\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   120
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   121
theorem bury_correct:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   122
  "(c,s) \<Rightarrow> s'  \<Longrightarrow> s = t on L c X \<Longrightarrow>
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   123
  \<exists> t'. (bury c X,t) \<Rightarrow> t' & s' = t' on X"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   124
proof (induction arbitrary: X t rule: big_step_induct)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   125
  case Skip then show ?case by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   126
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   127
  case Assign then show ?case
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   128
    by (auto simp: ball_Un)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   129
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   130
  case (Seq c1 s1 s2 c2 s3 X t1)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   131
  from Seq.IH(1) Seq.prems obtain t2 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   132
    t12: "(bury c1 (L c2 X), t1) \<Rightarrow> t2" and s2t2: "s2 = t2 on L c2 X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   133
    by simp blast
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   134
  from Seq.IH(2)[OF s2t2] obtain t3 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   135
    t23: "(bury c2 X, t2) \<Rightarrow> t3" and s3t3: "s3 = t3 on X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   136
    by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   137
  show ?case using t12 t23 s3t3 by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   138
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   139
  case (IfTrue b s c1 s' c2)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   140
  hence "s = t on vars b" "s = t on L c1 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   141
  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: 58889
diff changeset
   142
  from IfTrue.IH[OF \<open>s = t on L c1 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   143
    "(bury c1 X, t) \<Rightarrow> t'" "s' =t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   144
  thus ?case using \<open>bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   145
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   146
  case (IfFalse b s c2 s' c1)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   147
  hence "s = t on vars b" "s = t on L c2 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   148
  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: 58889
diff changeset
   149
  from IfFalse.IH[OF \<open>s = t on L c2 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   150
    "(bury c2 X, t) \<Rightarrow> t'" "s' = t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   151
  thus ?case using \<open>~bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   152
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   153
  case (WhileFalse b s c)
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   154
  hence "~ bval b t" by (metis L_While_vars bval_eq_if_eq_on_vars subsetD)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   155
  thus ?case
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   156
    by simp (metis L_While_X WhileFalse.prems big_step.WhileFalse subsetD)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   157
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   158
  case (WhileTrue b s1 c s2 s3 X t1)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   159
  let ?w = "WHILE b DO c"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   160
  from \<open>bval b s1\<close> WhileTrue.prems have "bval b t1"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   161
    by (metis L_While_vars bval_eq_if_eq_on_vars subsetD)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   162
  have "s1 = t1 on L c (L ?w X)"
45771
a70465244096 added lemmas
nipkow
parents: 45770
diff changeset
   163
    using L_While_pfp WhileTrue.prems by blast
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   164
  from WhileTrue.IH(1)[OF this] obtain t2 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   165
    "(bury c (L ?w X), t1) \<Rightarrow> t2" "s2 = t2 on L ?w X" by auto
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   166
  from WhileTrue.IH(2)[OF this(2)] obtain t3
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   167
    where "(bury ?w X,t2) \<Rightarrow> t3" "s3 = t3 on X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   168
    by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   169
  with \<open>bval b t1\<close> \<open>(bury c (L ?w X), t1) \<Rightarrow> t2\<close> show ?case by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   170
qed
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   171
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   172
corollary final_bury_correct: "(c,s) \<Rightarrow> s' \<Longrightarrow> (bury c UNIV,s) \<Rightarrow> s'"
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   173
using bury_correct[of c s s' UNIV]
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   174
by (auto simp: fun_eq_iff[symmetric])
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   175
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   176
text\<open>Now the opposite direction.\<close>
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   177
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   178
lemma SKIP_bury[simp]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   179
  "SKIP = bury c X \<longleftrightarrow> c = SKIP | (\<exists>x a. c = x::=a & x \<notin> X)"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   180
by (cases c) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   181
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   182
lemma Assign_bury[simp]: "x::=a = bury c X \<longleftrightarrow> c = x::=a \<and> x \<in> X"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   183
by (cases c) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   184
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52046
diff changeset
   185
lemma Seq_bury[simp]: "bc\<^sub>1;;bc\<^sub>2 = bury c X \<longleftrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   186
  (\<exists>c\<^sub>1 c\<^sub>2. c = c\<^sub>1;;c\<^sub>2 & bc\<^sub>2 = bury c\<^sub>2 X & bc\<^sub>1 = bury c\<^sub>1 (L c\<^sub>2 X))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   187
by (cases c) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   188
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   189
lemma If_bury[simp]: "IF b THEN bc1 ELSE bc2 = bury c X \<longleftrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   190
  (\<exists>c1 c2. c = IF b THEN c1 ELSE c2 &
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   191
     bc1 = bury c1 X & bc2 = bury c2 X)"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   192
by (cases c) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   193
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   194
lemma While_bury[simp]: "WHILE b DO bc' = bury c X \<longleftrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   195
  (\<exists>c'. c = WHILE b DO c' & bc' = bury c' (L (WHILE b DO c') X))"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   196
by (cases c) auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   197
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   198
theorem bury_correct2:
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   199
  "(bury c X,s) \<Rightarrow> s'  \<Longrightarrow> s = t on L c X \<Longrightarrow>
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   200
  \<exists> t'. (c,t) \<Rightarrow> t' & s' = t' on X"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   201
proof (induction "bury c X" s s' arbitrary: c X t rule: big_step_induct)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   202
  case Skip then show ?case by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   203
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   204
  case Assign then show ?case
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   205
    by (auto simp: ball_Un)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   206
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   207
  case (Seq bc1 s1 s2 bc2 s3 c X t1)
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51975
diff changeset
   208
  then obtain c1 c2 where c: "c = c1;;c2"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   209
    and bc2: "bc2 = bury c2 X" and bc1: "bc1 = bury c1 (L c2 X)" by auto
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   210
  note IH = Seq.hyps(2,4)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45784
diff changeset
   211
  from IH(1)[OF bc1, of t1] Seq.prems c obtain t2 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   212
    t12: "(c1, t1) \<Rightarrow> t2" and s2t2: "s2 = t2 on L c2 X" by auto
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   213
  from IH(2)[OF bc2 s2t2] obtain t3 where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   214
    t23: "(c2, t2) \<Rightarrow> t3" and s3t3: "s3 = t3 on X"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   215
    by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   216
  show ?case using c t12 t23 s3t3 by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   217
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   218
  case (IfTrue b s bc1 s' bc2)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   219
  then obtain c1 c2 where c: "c = IF b THEN c1 ELSE c2"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   220
    and bc1: "bc1 = bury c1 X" and bc2: "bc2 = bury c2 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   221
  have "s = t on vars b" "s = t on L c1 X" using IfTrue.prems c by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   222
  from bval_eq_if_eq_on_vars[OF this(1)] IfTrue(1) have "bval b t" by simp
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   223
  note IH = IfTrue.hyps(3)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   224
  from IH[OF bc1 \<open>s = t on L c1 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   225
    "(c1, t) \<Rightarrow> t'" "s' =t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   226
  thus ?case using c \<open>bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   227
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   228
  case (IfFalse b s bc2 s' bc1)
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   229
  then obtain c1 c2 where c: "c = IF b THEN c1 ELSE c2"
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   230
    and bc1: "bc1 = bury c1 X" and bc2: "bc2 = bury c2 X" by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   231
  have "s = t on vars b" "s = t on L c2 X" using IfFalse.prems c by auto
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   232
  from bval_eq_if_eq_on_vars[OF this(1)] IfFalse(1) have "~bval b t" by simp
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   233
  note IH = IfFalse.hyps(3)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   234
  from IH[OF bc2 \<open>s = t on L c2 X\<close>] obtain t' where
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   235
    "(c2, t) \<Rightarrow> t'" "s' =t' on X" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   236
  thus ?case using c \<open>~bval b t\<close> by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   237
next
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   238
  case (WhileFalse b s c)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   239
  hence "~ bval b t"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   240
    by auto (metis L_While_vars bval_eq_if_eq_on_vars rev_subsetD)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   241
  thus ?case using WhileFalse
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   242
    by auto (metis L_While_X big_step.WhileFalse subsetD)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   243
next
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   244
  case (WhileTrue b s1 bc' s2 s3 w X t1)
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   245
  then obtain c' where w: "w = WHILE b DO c'"
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   246
    and bc': "bc' = bury c' (L (WHILE b DO c') X)" by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   247
  from \<open>bval b s1\<close> WhileTrue.prems w have "bval b t1"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   248
    by auto (metis L_While_vars bval_eq_if_eq_on_vars subsetD)
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44923
diff changeset
   249
  note IH = WhileTrue.hyps(3,5)
51468
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   250
  have "s1 = t1 on L c' (L w X)"
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   251
    using L_While_pfp WhileTrue.prems w by blast
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   252
  with IH(1)[OF bc', of t1] w obtain t2 where
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   253
    "(c', t1) \<Rightarrow> t2" "s2 = t2 on L w X" by auto
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   254
  from IH(2)[OF WhileTrue.hyps(6), of t2] w this(2) obtain t3
0e8012983c4e proofs depend only on constraints, not on def of L WHILE
nipkow
parents: 51448
diff changeset
   255
    where "(w,t2) \<Rightarrow> t3" "s3 = t3 on X"
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   256
    by auto
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58889
diff changeset
   257
  with \<open>bval b t1\<close> \<open>(c', t1) \<Rightarrow> t2\<close> w show ?case by auto
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   258
qed
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   259
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   260
corollary final_bury_correct2: "(bury c UNIV,s) \<Rightarrow> s' \<Longrightarrow> (c,s) \<Rightarrow> s'"
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   261
using bury_correct2[of c UNIV]
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   262
by (auto simp: fun_eq_iff[symmetric])
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   263
51433
nipkow
parents: 51425
diff changeset
   264
corollary bury_sim: "bury c UNIV \<sim> c"
51975
7bab3fb52e3e tuned names
nipkow
parents: 51468
diff changeset
   265
by(metis final_bury_correct final_bury_correct2)
43158
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   266
686fa0a0696e imported rest of new IMP
kleing
parents:
diff changeset
   267
end