src/HOL/IMP/Fold.thy
author blanchet
Wed, 24 Sep 2014 15:45:55 +0200
changeset 58425 246985c6b20b
parent 55583 a0134252ac29
child 67613 ce654b0e6d69
permissions -rw-r--r--
simpler proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
     1
theory Fold imports Sem_Equiv Vars begin
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     2
44850
a6095c96a89b tuned headers
nipkow
parents: 44070
diff changeset
     3
subsection "Simple folding of arithmetic expressions"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     4
45134
9b02f6665fc8 discontinued obsolete 'types' command;
wenzelm
parents: 45015
diff changeset
     5
type_synonym
45212
e87feee00a4c renamed name -> vname
nipkow
parents: 45200
diff changeset
     6
  tab = "vname \<Rightarrow> val option"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     7
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
     8
fun afold :: "aexp \<Rightarrow> tab \<Rightarrow> aexp" where
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
     9
"afold (N n) _ = N n" |
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    10
"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
    11
"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
    12
  (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
    13
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    14
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
    15
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    16
theorem aval_afold[simp]:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    17
assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    18
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
    19
  using assms
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    20
  by (induct a) (auto simp: approx_def split: aexp.split option.split)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    21
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    22
theorem aval_afold_N:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    23
assumes "approx t s"
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    24
shows "afold a t = N n \<Longrightarrow> aval a s = n"
51573
acc4bd79e2e9 replace induction by hammer
kleing
parents: 51514
diff changeset
    25
  by (metis assms aval.simps(1) aval_afold)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    26
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    27
definition
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    28
  "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
    29
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    30
primrec "defs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    31
"defs SKIP t = t" |
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    32
"defs (x ::= a) t =
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    33
  (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
    34
"defs (c1;;c2) t = (defs c2 o defs c1) t" |
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    35
"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
    36
"defs (WHILE b DO c) t = t |` (-lvars c)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    37
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    38
primrec fold where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    39
"fold SKIP _ = SKIP" |
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
    40
"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
    41
"fold (c1;;c2) t = (fold c1 t;; fold c2 (defs c1 t))" |
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    42
"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
    43
"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
    44
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    45
lemma approx_merge:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    46
  "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
    47
  by (fastforce simp: merge_def approx_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    48
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    49
lemma approx_map_le:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    50
  "approx t2 s \<Longrightarrow> t1 \<subseteq>\<^sub>m t2 \<Longrightarrow> approx t1 s"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    51
  by (clarsimp simp: approx_def map_le_def dom_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    52
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    53
lemma restrict_map_le [intro!, simp]: "t |` S \<subseteq>\<^sub>m t"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    54
  by (clarsimp simp: restrict_map_def map_le_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    55
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    56
lemma merge_restrict:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    57
  assumes "t1 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    58
  assumes "t2 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    59
  shows "merge t1 t2 |` S = t |` S"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    60
proof -
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    61
  from assms
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    62
  have "\<forall>x. (t1 |` S) x = (t |` S) x"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    63
   and "\<forall>x. (t2 |` S) x = (t |` S) x" by auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    64
  thus ?thesis
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    65
    by (auto simp: merge_def restrict_map_def
51506
cd573f1a82b2 removed obsolete uses of ext
kleing
parents: 48909
diff changeset
    66
             split: if_splits)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    67
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    68
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    69
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    70
lemma defs_restrict:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    71
  "defs c t |` (- lvars c) = t |` (- lvars c)"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
    72
proof (induction c arbitrary: t)
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
    73
  case (Seq c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    74
  hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    75
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    76
  hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    77
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    78
  moreover
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
    79
  from Seq
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    80
  have "defs c2 (defs c1 t) |` (- lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    81
        defs c1 t |` (- lvars c2)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    82
    by simp
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    83
  hence "defs c2 (defs c1 t) |` (- lvars c2) |` (- lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    84
         defs c1 t |` (- lvars c2) |` (- lvars c1)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    85
    by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    86
  ultimately
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    87
  show ?case by (clarsimp simp: Int_commute)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    88
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    89
  case (If b c1 c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    90
  hence "defs c1 t |` (- lvars c1) = t |` (- lvars c1)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    91
  hence "defs c1 t |` (- lvars c1) |` (-lvars c2) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    92
         t |` (- lvars c1) |` (-lvars c2)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    93
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    94
  from If
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    95
  have "defs c2 t |` (- lvars c2) = t |` (- lvars c2)" by simp
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    96
  hence "defs c2 t |` (- lvars c2) |` (-lvars c1) =
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
    97
         t |` (- lvars c2) |` (-lvars c1)" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    98
  ultimately
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    99
  show ?case by (auto simp: Int_commute intro: merge_restrict)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   100
qed (auto split: aexp.split)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   101
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   102
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   103
lemma big_step_pres_approx:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   104
  "(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
   105
proof (induction arbitrary: t rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   106
  case Skip thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   107
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   108
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   109
  thus ?case
51514
9bed72e55475 simp_const -> afold; bfold -> fold'; bsimp_const -> bfold
kleing
parents: 51506
diff changeset
   110
    by (clarsimp simp: aval_afold_N approx_def split: aexp.split)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   111
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   112
  case (Seq c1 s1 s2 c2 s3)
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   113
  have "approx (defs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems])
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   114
  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
   115
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   116
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   117
  case (IfTrue b s c1 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   118
  hence "approx (defs c1 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   119
  thus ?case by (simp add: approx_merge)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   120
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   121
  case (IfFalse b s c2 s')
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   122
  hence "approx (defs c2 t) s'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   123
  thus ?case by (simp add: approx_merge)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   124
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   125
  case WhileFalse
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   126
  thus ?case by (simp add: approx_def restrict_map_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   127
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   128
  case (WhileTrue b s1 c s2 s3)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   129
  hence "approx (defs c t) s2" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   130
  with WhileTrue
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   131
  have "approx (defs c t |` (-lvars c)) s3" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   132
  thus ?case by (simp add: defs_restrict)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   133
qed
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   134
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   135
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   136
lemma big_step_pres_approx_restrict:
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   137
  "(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
   138
proof (induction arbitrary: t rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   139
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   140
  thus ?case by (clarsimp simp: approx_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   141
next
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   142
  case (Seq c1 s1 s2 c2 s3)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   143
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s1"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   144
    by (simp add: Int_commute)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   145
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s2"
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   146
    by (rule Seq)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   147
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s2"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   148
    by (simp add: Int_commute)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   149
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s3"
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45212
diff changeset
   150
    by (rule Seq)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   151
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   152
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   153
  case (IfTrue b s c1 s' c2)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   154
  hence "approx (t |` (-lvars c2) |` (-lvars c1)) s"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   155
    by (simp add: Int_commute)
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 (rule IfTrue)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   158
  thus ?case by (simp add: Int_commute)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   159
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   160
  case (IfFalse b s c2 s' c1)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   161
  hence "approx (t |` (-lvars c1) |` (-lvars c2)) s"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   162
    by simp
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 (rule IfFalse)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   165
  thus ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   166
qed auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   167
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   168
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   169
declare assign_simp [simp]
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   170
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   171
lemma approx_eq:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   172
  "approx t \<Turnstile> c \<sim> fold c t"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
   173
proof (induction c arbitrary: t)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   174
  case SKIP show ?case by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   175
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   176
  case Assign
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   177
  show ?case by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   178
next
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   179
  case Seq
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   180
  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
   181
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   182
  case If
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   183
  thus ?case by (auto intro!: equiv_up_to_if_weak)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   184
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   185
  case (While b c)
52072
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   186
  hence "approx (t |` (- lvars c)) \<Turnstile>
c25764d7246e defined lvars and rvars of commands separately.
nipkow
parents: 52046
diff changeset
   187
         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
   188
    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
   189
  thus ?case
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   190
    by (auto intro: equiv_up_to_weaken approx_map_le)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   191
qed
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   192
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   193
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   194
lemma approx_empty [simp]:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   195
  "approx empty = (\<lambda>_. True)"
51506
cd573f1a82b2 removed obsolete uses of ext
kleing
parents: 48909
diff changeset
   196
  by (auto simp: approx_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   197
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   198
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   199
theorem constant_folding_equiv:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   200
  "fold c empty \<sim> c"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   201
  using approx_eq [of empty c]
52825
71fef62c4213 removed duplicate lemma
kleing
parents: 52072
diff changeset
   202
  by (simp add: equiv_up_to_True sim_sym)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   203
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   204
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   205
end