src/HOL/IMP/Sem_Equiv.thy
author wenzelm
Tue, 04 Nov 2014 17:37:15 +0100
changeset 58896 5a2f475e2ded
parent 58889 5b7a9633cfa8
child 67406 23307fd33906
permissions -rw-r--r--
approximative update of versions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 55598
diff changeset
     1
section "Constant Folding"
54297
nipkow
parents: 54296
diff changeset
     2
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     3
theory Sem_Equiv
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
     4
imports Big_Step
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     5
begin
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     6
54296
nipkow
parents: 52121
diff changeset
     7
subsection "Semantic Equivalence up to a Condition"
nipkow
parents: 52121
diff changeset
     8
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
     9
type_synonym assn = "state \<Rightarrow> bool"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    10
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    11
definition
52021
59963cda805a explicitly state equivalence relation for sim; tweak syntax of sem_equiv
kleing
parents: 48909
diff changeset
    12
  equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" ("_ \<Turnstile> _ \<sim> _" [50,0,10] 50)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    13
where
52121
5b889b1b465b prefer object equality
kleing
parents: 52046
diff changeset
    14
  "(P \<Turnstile> c \<sim> c') = (\<forall>s s'. P s \<longrightarrow> (c,s) \<Rightarrow> s' \<longleftrightarrow> (c',s) \<Rightarrow> s')"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    15
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    16
definition
52021
59963cda805a explicitly state equivalence relation for sim; tweak syntax of sem_equiv
kleing
parents: 48909
diff changeset
    17
  bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" ("_ \<Turnstile> _ <\<sim>> _" [50,0,10] 50)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    18
where
52121
5b889b1b465b prefer object equality
kleing
parents: 52046
diff changeset
    19
  "(P \<Turnstile> b <\<sim>> b') = (\<forall>s. P s \<longrightarrow> bval b s = bval b' s)"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    20
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    21
lemma equiv_up_to_True:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    22
  "((\<lambda>_. True) \<Turnstile> c \<sim> c') = (c \<sim> c')"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    23
  by (simp add: equiv_def equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    24
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    25
lemma equiv_up_to_weaken:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    26
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> (\<And>s. P' s \<Longrightarrow> P s) \<Longrightarrow> P' \<Turnstile> c \<sim> c'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    27
  by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    28
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    29
lemma equiv_up_toI:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    30
  "(\<And>s s'. P s \<Longrightarrow> (c, s) \<Rightarrow> s' = (c', s) \<Rightarrow> s') \<Longrightarrow> P \<Turnstile> c \<sim> c'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    31
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    32
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    33
lemma equiv_up_toD1:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    34
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> (c, s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> (c', s) \<Rightarrow> s'"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    35
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    36
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    37
lemma equiv_up_toD2:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    38
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> (c', s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> (c, s) \<Rightarrow> s'"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    39
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    40
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    41
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    42
lemma equiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    43
  "P \<Turnstile> c \<sim> c"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    44
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    45
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    46
lemma equiv_up_to_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    47
  "(P \<Turnstile> c \<sim> c') = (P \<Turnstile> c' \<sim> c)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    48
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    49
45218
f115540543d8 removed [trans] concept from basic material
kleing
parents: 45200
diff changeset
    50
lemma equiv_up_to_trans:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    51
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> P \<Turnstile> c' \<sim> c'' \<Longrightarrow> P \<Turnstile> c \<sim> c''"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    52
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    53
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 bequiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    56
  "P \<Turnstile> b <\<sim>> b"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    57
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    58
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    59
lemma bequiv_up_to_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    60
  "(P \<Turnstile> b <\<sim>> b') = (P \<Turnstile> b' <\<sim>> b)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    61
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    62
45218
f115540543d8 removed [trans] concept from basic material
kleing
parents: 45200
diff changeset
    63
lemma bequiv_up_to_trans:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    64
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> P \<Turnstile> b' <\<sim>> b'' \<Longrightarrow> P \<Turnstile> b <\<sim>> b''"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    65
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    66
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    67
lemma bequiv_up_to_subst:
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    68
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> P s \<Longrightarrow> bval b s = bval b' s"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    69
  by (simp add: bequiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    70
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    71
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45218
diff changeset
    72
lemma equiv_up_to_seq:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    73
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> Q \<Turnstile> d \<sim> d' \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    74
  (\<And>s s'. (c,s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> Q s') \<Longrightarrow>
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 52021
diff changeset
    75
  P \<Turnstile> (c;; d) \<sim> (c';; d')"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    76
  by (clarsimp simp: equiv_up_to_def) blast
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    77
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
    78
lemma equiv_up_to_while_lemma_weak:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    79
  shows "(d,s) \<Rightarrow> s' \<Longrightarrow>
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    80
         P \<Turnstile> b <\<sim>> b' \<Longrightarrow>
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
    81
         P \<Turnstile> c \<sim> c' \<Longrightarrow>
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    82
         (\<And>s s'. (c, s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> bval b s \<Longrightarrow> P s') \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    83
         P s \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    84
         d = WHILE b DO c \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    85
         (WHILE b' DO c', s) \<Rightarrow> s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
    86
proof (induction rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    87
  case (WhileTrue b s1 c s2 s3)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    88
  hence IH: "P s2 \<Longrightarrow> (WHILE b' DO c', s2) \<Rightarrow> s3" by auto
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    89
  from WhileTrue.prems
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    90
  have "P \<Turnstile> b <\<sim>> b'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    91
  with `bval b s1` `P s1`
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    92
  have "bval b' s1" by (simp add: bequiv_up_to_def)
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 WhileTrue.prems
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
    95
  have "P \<Turnstile> c \<sim> c'" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    96
  with `bval b s1` `P s1` `(c, s1) \<Rightarrow> s2`
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    97
  have "(c', s1) \<Rightarrow> s2" by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    98
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    99
  from WhileTrue.prems
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   100
  have "\<And>s s'. (c,s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> bval b s \<Longrightarrow> P s'" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   101
  with `P s1` `bval b s1` `(c, s1) \<Rightarrow> s2`
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   102
  have "P s2" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   103
  hence "(WHILE b' DO c', s2) \<Rightarrow> s3" by (rule IH)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   104
  ultimately
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   105
  show ?case by blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   106
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   107
  case WhileFalse
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   108
  thus ?case by (auto simp: bequiv_up_to_def)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   109
qed (fastforce simp: equiv_up_to_def bequiv_up_to_def)+
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   110
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   111
lemma equiv_up_to_while_weak:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   112
  assumes b: "P \<Turnstile> b <\<sim>> b'"
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   113
  assumes c: "P \<Turnstile> c \<sim> c'"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   114
  assumes I: "\<And>s s'. (c, s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> bval b s \<Longrightarrow> P s'"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   115
  shows "P \<Turnstile> WHILE b DO c \<sim> WHILE b' DO c'"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   116
proof -
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   117
  from b have b': "P \<Turnstile> b' <\<sim>> b" by (simp add: bequiv_up_to_sym)
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   118
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   119
  from c b have c': "P \<Turnstile> c' \<sim> c" by (simp add: equiv_up_to_sym)
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   120
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   121
  from I
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   122
  have I': "\<And>s s'. (c', s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> bval b' s \<Longrightarrow> P s'"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   123
    by (auto dest!: equiv_up_toD1 [OF c'] simp: bequiv_up_to_subst [OF b'])
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   124
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   125
  note equiv_up_to_while_lemma_weak [OF _ b c]
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   126
       equiv_up_to_while_lemma_weak [OF _ b' c']
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   127
  thus ?thesis using I I' by (auto intro!: equiv_up_toI)
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   128
qed
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   129
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   130
lemma equiv_up_to_if_weak:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   131
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> P \<Turnstile> c \<sim> c' \<Longrightarrow> P \<Turnstile> d \<sim> d' \<Longrightarrow>
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   132
   P \<Turnstile> IF b THEN c ELSE d \<sim> IF b' THEN c' ELSE d'"
55598
da35747597bd move stronger rules into exercise
kleing
parents: 54297
diff changeset
   133
  by (auto simp: bequiv_up_to_def equiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   134
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   135
lemma equiv_up_to_if_True [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   136
  "(\<And>s. P s \<Longrightarrow> bval b s) \<Longrightarrow> P \<Turnstile> IF b THEN c1 ELSE c2 \<sim> c1"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   137
  by (auto simp: equiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   138
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   139
lemma equiv_up_to_if_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   140
  "(\<And>s. P s \<Longrightarrow> \<not> bval b s) \<Longrightarrow> P \<Turnstile> IF b THEN c1 ELSE c2 \<sim> c2"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   141
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   142
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   143
lemma equiv_up_to_while_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   144
  "(\<And>s. P s \<Longrightarrow> \<not> bval b s) \<Longrightarrow> P \<Turnstile> WHILE b DO c \<sim> SKIP"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   145
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   146
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
   147
lemma while_never: "(c, s) \<Rightarrow> u \<Longrightarrow> c \<noteq> WHILE (Bc True) DO c'"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   148
 by (induct rule: big_step_induct) auto
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   149
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   150
lemma equiv_up_to_while_True [intro!,simp]:
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
   151
  "P \<Turnstile> WHILE Bc True DO c \<sim> WHILE Bc True DO SKIP"
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   152
  unfolding equiv_up_to_def
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   153
  by (blast dest: while_never)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   154
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   155
44261
e44f465c00a1 HOL-IMP: respect set/pred distinction
huffman
parents: 44070
diff changeset
   156
end