src/HOL/IMP/Fold.thy
author traytel
Thu, 08 Aug 2013 15:30:25 +0200
changeset 52912 bdd610910e2c
parent 52825 71fef62c4213
child 54297 3fc1b77ef750
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     1
header "Constant Folding"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     2
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
     3
theory Fold imports Sem_Equiv Vars begin
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     4
44850
a6095c96a89b tuned headers
nipkow
parents: 44070
diff changeset
     5
subsection "Simple folding of arithmetic expressions"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     6
45134
9b02f6665fc8 discontinued obsolete 'types' command;
wenzelm
parents: 45015
diff changeset
     7
type_synonym
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
     8
  tab = "vname \<Rightarrow> val option"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     9
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    10
fun afold :: "aexp \<Rightarrow> tab \<Rightarrow> aexp" where
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    11
"afold (N n) _ = N n" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    12
"afold (V x) t = (case t x of None \<Rightarrow> V x | Some k \<Rightarrow> N k)" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    13
"afold (Plus e1 e2) t = (case (afold e1 t, afold e2 t) of
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    14
  (N n1, N n2) \<Rightarrow> N(n1+n2) | (e1',e2') \<Rightarrow> Plus e1' e2')"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    15
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    16
definition "approx t s \<longleftrightarrow> (ALL x k. t x = Some k \<longrightarrow> s x = k)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    17
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    18
theorem aval_afold[simp]:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    19
assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    20
shows "aval (afold a t) s = aval a s"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    21
  using assms
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    22
  by (induct a) (auto simp: approx_def split: aexp.split option.split)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    23
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    24
theorem aval_afold_N:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    25
assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    26
shows "afold a t = N n \<Longrightarrow> aval a s = n"
51573
acc4bd79e2e9 replace induction by hammer
kleing
parents: 51514
diff changeset
    27
  by (metis assms aval.simps(1) aval_afold)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    28
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    29
definition
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    30
  "merge t1 t2 = (\<lambda>m. if t1 m = t2 m then t1 m else None)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    31
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    32
primrec "defs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    33
"defs SKIP t = t" |
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    34
"defs (x ::= a) t =
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    35
  (case afold a t of N k \<Rightarrow> t(x \<mapsto> k) | _ \<Rightarrow> t(x:=None))" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51573
diff changeset
    36
"defs (c1;;c2) t = (defs c2 o defs c1) t" |
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    37
"defs (IF b THEN c1 ELSE c2) t = merge (defs c1 t) (defs c2 t)" |
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    38
"defs (WHILE b DO c) t = t |` (-lvars c)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    39
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    40
primrec fold where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    41
"fold SKIP _ = SKIP" |
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    42
"fold (x ::= a) t = (x ::= (afold a t))" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51573
diff changeset
    43
"fold (c1;;c2) t = (fold c1 t;; fold c2 (defs c1 t))" |
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    44
"fold (IF b THEN c1 ELSE c2) t = IF b THEN fold c1 t ELSE fold c2 t" |
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    45
"fold (WHILE b DO c) t = WHILE b DO fold c (t |` (-lvars c))"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    46
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    47
lemma approx_merge:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    48
  "approx t1 s \<or> approx t2 s \<Longrightarrow> approx (merge t1 t2) s"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44850
diff changeset
    49
  by (fastforce simp: merge_def approx_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    50
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    51
lemma approx_map_le:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    52
  "approx t2 s \<Longrightarrow> t1 \<subseteq>\<^sub>m t2 \<Longrightarrow> approx t1 s"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    53
  by (clarsimp simp: approx_def map_le_def dom_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    54
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    55
lemma restrict_map_le [intro!, simp]: "t |` S \<subseteq>\<^sub>m t"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    56
  by (clarsimp simp: restrict_map_def map_le_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    57
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    58
lemma merge_restrict:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    59
  assumes "t1 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    60
  assumes "t2 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    61
  shows "merge t1 t2 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    62
proof -
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    63
  from assms
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    64
  have "\<forall>x. (t1 |` S) x = (t |` S) x"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    65
   and "\<forall>x. (t2 |` S) x = (t |` S) x" by auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    66
  thus ?thesis
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    67
    by (auto simp: merge_def restrict_map_def
51506
cd573f1a82b2 removed obsolete uses of ext
kleing
parents: 48909
diff changeset
    68
             split: if_splits)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    69
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    70
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    71
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    72
lemma defs_restrict:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    73
  "defs c t |` (- lvars c) = t |` (- lvars c)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
    74
proof (induction c arbitrary: t)
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
    75
  case (Seq c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    76
  hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    77
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    78
  hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    79
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    80
  moreover
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
    81
  from Seq
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    82
  have "defs c2 (defs c1 t) |` (- lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    83
        defs c1 t |` (- lvars c2)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    84
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    85
  hence "defs c2 (defs c1 t) |` (- lvars c2) |` (- lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    86
         defs c1 t |` (- lvars c2) |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    87
    by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    88
  ultimately
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    89
  show ?case by (clarsimp simp: Int_commute)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    90
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    91
  case (If b c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    92
  hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    93
  hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    94
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    95
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    96
  from If
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    97
  have "defs c2 t |` (- lvars c2) = t |` (- lvars c2)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    98
  hence "defs c2 t |` (- lvars c2) |` (-lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    99
         t |` (- lvars c2) |` (-lvars c1)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   100
  ultimately
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   101
  show ?case by (auto simp: Int_commute intro: merge_restrict)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   102
qed (auto split: aexp.split)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   103
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   104
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   105
lemma big_step_pres_approx:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   106
  "(c,s) \<Rightarrow> s' \<Longrightarrow> approx t s \<Longrightarrow> approx (defs c t) s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   107
proof (induction arbitrary: t rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   108
  case Skip thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   109
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   110
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   111
  thus ?case
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   112
    by (clarsimp simp: aval_afold_N approx_def split: aexp.split)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   113
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   114
  case (Seq c1 s1 s2 c2 s3)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   115
  have "approx (defs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems])
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   116
  hence "approx (defs c2 (defs c1 t)) s3" by (rule Seq.IH(2))
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   117
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   118
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   119
  case (IfTrue b s c1 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   120
  hence "approx (defs c1 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   121
  thus ?case by (simp add: approx_merge)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   122
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   123
  case (IfFalse b s c2 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   124
  hence "approx (defs c2 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   125
  thus ?case by (simp add: approx_merge)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   126
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   127
  case WhileFalse
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   128
  thus ?case by (simp add: approx_def restrict_map_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   129
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   130
  case (WhileTrue b s1 c s2 s3)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   131
  hence "approx (defs c t) s2" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   132
  with WhileTrue
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   133
  have "approx (defs c t |` (-lvars c)) s3" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   134
  thus ?case by (simp add: defs_restrict)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   135
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   136
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   137
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   138
lemma big_step_pres_approx_restrict:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   139
  "(c,s) \<Rightarrow> s' \<Longrightarrow> approx (t |` (-lvars c)) s \<Longrightarrow> approx (t |` (-lvars c)) s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   140
proof (induction arbitrary: t rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   141
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   142
  thus ?case by (clarsimp simp: approx_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   143
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   144
  case (Seq c1 s1 s2 c2 s3)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   145
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s1"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   146
    by (simp add: Int_commute)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   147
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s2"
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   148
    by (rule Seq)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   149
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s2"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   150
    by (simp add: Int_commute)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   151
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s3"
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   152
    by (rule Seq)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   153
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   154
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   155
  case (IfTrue b s c1 s' c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   156
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   157
    by (simp add: Int_commute)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   158
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s'"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   159
    by (rule IfTrue)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   160
  thus ?case by (simp add: Int_commute)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   161
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   162
  case (IfFalse b s c2 s' c1)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   163
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   164
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   165
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s'"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   166
    by (rule IfFalse)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   167
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   168
qed auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   169
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   170
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   171
declare assign_simp [simp]
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   172
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   173
lemma approx_eq:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   174
  "approx t \<Turnstile> c \<sim> fold c t"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   175
proof (induction c arbitrary: t)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   176
  case SKIP show ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   177
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   178
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   179
  show ?case by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   180
next
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   181
  case Seq
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   182
  thus ?case by (auto intro!: equiv_up_to_seq big_step_pres_approx)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   183
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   184
  case If
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   185
  thus ?case by (auto intro!: equiv_up_to_if_weak)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   186
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   187
  case (While b c)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   188
  hence "approx (t |` (- lvars c)) \<Turnstile>
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   189
         WHILE b DO c \<sim> WHILE b DO fold c (t |` (- lvars c))"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   190
    by (auto intro: equiv_up_to_while_weak big_step_pres_approx_restrict)
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   191
  thus ?case
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   192
    by (auto intro: equiv_up_to_weaken approx_map_le)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   193
qed
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   194
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   195
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   196
lemma approx_empty [simp]:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   197
  "approx empty = (\<lambda>_. True)"
51506
cd573f1a82b2 removed obsolete uses of ext
kleing
parents: 48909
diff changeset
   198
  by (auto simp: approx_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   199
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   200
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   201
theorem constant_folding_equiv:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   202
  "fold c empty \<sim> c"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   203
  using approx_eq [of empty c]
52825
71fef62c4213 removed duplicate lemma
kleing
parents: 52072
diff changeset
   204
  by (simp add: equiv_up_to_True sim_sym)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   205
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   206
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   207
44850
a6095c96a89b tuned headers
nipkow
parents: 44070
diff changeset
   208
subsection {* More ambitious folding including boolean expressions *}
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   209
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   210
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   211
fun bfold :: "bexp \<Rightarrow> tab \<Rightarrow> bexp" where
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   212
"bfold (Less a1 a2) t = less (afold a1 t) (afold a2 t)" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   213
"bfold (And b1 b2) t = and (bfold b1 t) (bfold b2 t)" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   214
"bfold (Not b) t = not(bfold b t)" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   215
"bfold (Bc bc) _ = Bc bc"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   216
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   217
theorem bval_bfold [simp]:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   218
  assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   219
  shows "bval (bfold b t) s = bval b s"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   220
  using assms by (induct b) auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   221
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   222
lemma not_Bc [simp]: "not (Bc v) = Bc (\<not>v)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   223
  by (cases v) auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   224
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   225
lemma not_Bc_eq [simp]: "(not b = Bc v) = (b = Bc (\<not>v))"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   226
  by (cases b) auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   227
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   228
lemma and_Bc_eq:
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   229
  "(and a b = Bc v) =
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   230
   (a = Bc False \<and> \<not>v  \<or>  b = Bc False \<and> \<not>v \<or>
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   231
    (\<exists>v1 v2. a = Bc v1 \<and> b = Bc v2 \<and> v = v1 \<and> v2))"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   232
  by (rule and.induct) auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   233
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   234
lemma less_Bc_eq [simp]:
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   235
  "(less a b = Bc v) = (\<exists>n1 n2. a = N n1 \<and> b = N n2 \<and> v = (n1 < n2))"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   236
  by (rule less.induct) auto
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   237
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   238
theorem bval_bfold_Bc:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   239
assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   240
shows "bfold b t = Bc v \<Longrightarrow> bval b s = v"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   241
  using assms
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   242
  by (induct b arbitrary: v)
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   243
     (auto simp: approx_def and_Bc_eq aval_afold_N
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   244
           split: bexp.splits bool.splits)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   245
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   246
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   247
primrec "bdefs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   248
"bdefs SKIP t = t" |
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   249
"bdefs (x ::= a) t =
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   250
  (case afold a t of N k \<Rightarrow> t(x \<mapsto> k) | _ \<Rightarrow> t(x:=None))" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51573
diff changeset
   251
"bdefs (c1;;c2) t = (bdefs c2 o bdefs c1) t" |
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   252
"bdefs (IF b THEN c1 ELSE c2) t = (case bfold b t of
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   253
    Bc True \<Rightarrow> bdefs c1 t
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   254
  | Bc False \<Rightarrow> bdefs c2 t
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   255
  | _ \<Rightarrow> merge (bdefs c1 t) (bdefs c2 t))" |
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   256
"bdefs (WHILE b DO c) t = t |` (-lvars c)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   257
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   258
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   259
primrec fold' where
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   260
"fold' SKIP _ = SKIP" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   261
"fold' (x ::= a) t = (x ::= (afold a t))" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 51573
diff changeset
   262
"fold' (c1;;c2) t = (fold' c1 t;; fold' c2 (bdefs c1 t))" |
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   263
"fold' (IF b THEN c1 ELSE c2) t = (case bfold b t of
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   264
    Bc True \<Rightarrow> fold' c1 t
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   265
  | Bc False \<Rightarrow> fold' c2 t
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   266
  | _ \<Rightarrow> IF bfold b t THEN fold' c1 t ELSE fold' c2 t)" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   267
"fold' (WHILE b DO c) t = (case bfold b t of
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45134
diff changeset
   268
    Bc False \<Rightarrow> SKIP
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   269
  | _ \<Rightarrow> WHILE bfold b (t |` (-lvars c)) DO fold' c (t |` (-lvars c)))"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   270
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   271
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   272
lemma bdefs_restrict:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   273
  "bdefs c t |` (- lvars c) = t |` (- lvars c)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   274
proof (induction c arbitrary: t)
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   275
  case (Seq c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   276
  hence "bdefs c1 t |` (- lvars c1) = t |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   277
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   278
  hence "bdefs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   279
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   280
  moreover
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   281
  from Seq
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   282
  have "bdefs c2 (bdefs c1 t) |` (- lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   283
        bdefs c1 t |` (- lvars c2)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   284
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   285
  hence "bdefs c2 (bdefs c1 t) |` (- lvars c2) |` (- lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   286
         bdefs c1 t |` (- lvars c2) |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   287
    by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   288
  ultimately
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   289
  show ?case by (clarsimp simp: Int_commute)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   290
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   291
  case (If b c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   292
  hence "bdefs c1 t |` (- lvars c1) = t |` (- lvars c1)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   293
  hence "bdefs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   294
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   295
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   296
  from If
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   297
  have "bdefs c2 t |` (- lvars c2) = t |` (- lvars c2)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   298
  hence "bdefs c2 t |` (- lvars c2) |` (-lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   299
         t |` (- lvars c2) |` (-lvars c1)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   300
  ultimately
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   301
  show ?case
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   302
    by (auto simp: Int_commute intro: merge_restrict
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   303
             split: bexp.splits bool.splits)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   304
qed (auto split: aexp.split bexp.split bool.split)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   305
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   306
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   307
lemma big_step_pres_approx_b:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   308
  "(c,s) \<Rightarrow> s' \<Longrightarrow> approx t s \<Longrightarrow> approx (bdefs c t) s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   309
proof (induction arbitrary: t rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   310
  case Skip thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   311
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   312
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   313
  thus ?case
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   314
    by (clarsimp simp: aval_afold_N approx_def split: aexp.split)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   315
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   316
  case (Seq c1 s1 s2 c2 s3)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   317
  have "approx (bdefs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems])
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   318
  hence "approx (bdefs c2 (bdefs c1 t)) s3" by (rule Seq.IH(2))
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   319
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   320
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   321
  case (IfTrue b s c1 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   322
  hence "approx (bdefs c1 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   323
  thus ?case using `bval b s` `approx t s`
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   324
    by (clarsimp simp: approx_merge bval_bfold_Bc
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   325
                 split: bexp.splits bool.splits)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   326
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   327
  case (IfFalse b s c2 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   328
  hence "approx (bdefs c2 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   329
  thus ?case using `\<not>bval b s` `approx t s`
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   330
    by (clarsimp simp: approx_merge bval_bfold_Bc
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   331
                 split: bexp.splits bool.splits)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   332
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   333
  case WhileFalse
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   334
  thus ?case
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   335
    by (clarsimp simp: bval_bfold_Bc approx_def restrict_map_def
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   336
                 split: bexp.splits bool.splits)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   337
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   338
  case (WhileTrue b s1 c s2 s3)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   339
  hence "approx (bdefs c t) s2" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   340
  with WhileTrue
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   341
  have "approx (bdefs c t |` (- lvars c)) s3"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   342
    by simp
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   343
  thus ?case
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   344
    by (simp add: bdefs_restrict)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   345
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   346
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   347
lemma fold'_equiv:
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   348
  "approx t \<Turnstile> c \<sim> fold' c t"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   349
proof (induction c arbitrary: t)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   350
  case SKIP show ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   351
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   352
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   353
  thus ?case by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   354
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   355
  case Seq
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   356
  thus ?case by (auto intro!: equiv_up_to_seq big_step_pres_approx_b)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   357
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   358
  case (If b c1 c2)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   359
  hence "approx t \<Turnstile> IF b THEN c1 ELSE c2 \<sim>
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   360
                   IF bfold b t THEN fold' c1 t ELSE fold' c2 t"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   361
    by (auto intro: equiv_up_to_if_weak simp: bequiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   362
  thus ?case using If
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   363
    by (fastforce simp: bval_bfold_Bc split: bexp.splits bool.splits
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   364
                 intro: equiv_up_to_trans)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   365
  next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   366
  case (While b c)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   367
  hence "approx (t |` (- lvars c)) \<Turnstile>
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   368
                   WHILE b DO c \<sim>
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   369
                   WHILE bfold b (t |` (- lvars c))
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   370
                      DO fold' c (t |` (- lvars c))" (is "_ \<Turnstile> ?W \<sim> ?W'")
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   371
    by (auto intro: equiv_up_to_while_weak big_step_pres_approx_restrict
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   372
             simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   373
  hence "approx t \<Turnstile> ?W \<sim> ?W'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   374
    by (auto intro: equiv_up_to_weaken approx_map_le)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   375
  thus ?case
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   376
    by (auto split: bexp.splits bool.splits
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   377
             intro: equiv_up_to_while_False
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   378
             simp: bval_bfold_Bc)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   379
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   380
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   381
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   382
theorem constant_folding_equiv':
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   383
  "fold' c empty \<sim> c"
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   384
  using fold'_equiv [of empty c]
52825
71fef62c4213 removed duplicate lemma
kleing
parents: 52072
diff changeset
   385
  by (simp add: equiv_up_to_True sim_sym)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   386
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   387
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   388
end