src/HOL/IMP/Sem_Equiv.thy
author nipkow
Tue, 20 Sep 2011 05:48:23 +0200
changeset 45015 fdac1e9880eb
parent 44890 22f665a2e91c
child 45200 1f1897ac7877
permissions -rw-r--r--
Updated IMP to use new induction method
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     1
header "Semantic Equivalence up to a Condition"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     2
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     3
theory Sem_Equiv
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     4
imports Hoare_Sound_Complete
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
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     7
definition
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     8
  equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" ("_ \<Turnstile> _ \<sim> _" [60,0,10] 60)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     9
where
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    10
  "P \<Turnstile> c \<sim> c' \<equiv> \<forall>s s'. P s \<longrightarrow> (c,s) \<Rightarrow> s' \<longleftrightarrow> (c',s) \<Rightarrow> s'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    11
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    12
definition 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    13
  bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" ("_ \<Turnstile> _ <\<sim>> _" [60,0,10] 60)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    14
where 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    15
  "P \<Turnstile> b <\<sim>> b' \<equiv> \<forall>s. P s \<longrightarrow> bval b s = bval b' s"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    16
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    17
lemma equiv_up_to_True:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    18
  "((\<lambda>_. True) \<Turnstile> c \<sim> c') = (c \<sim> c')"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    19
  by (simp add: equiv_def equiv_up_to_def)
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_weaken:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    22
  "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
    23
  by (simp add: 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_toI:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    26
  "(\<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
    27
  by (unfold equiv_up_to_def) blast
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_toD1:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    30
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> P s \<Longrightarrow> (c, s) \<Rightarrow> s' \<Longrightarrow> (c', s) \<Rightarrow> s'"
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_toD2:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    34
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> P s \<Longrightarrow> (c', s) \<Rightarrow> s' \<Longrightarrow> (c, s) \<Rightarrow> s'"
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
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    38
lemma equiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    39
  "P \<Turnstile> c \<sim> c"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    40
  by (auto simp: equiv_up_to_def)
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_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    43
  "(P \<Turnstile> c \<sim> c') = (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_trans [trans]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    47
  "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
    48
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    49
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 bequiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    52
  "P \<Turnstile> b <\<sim>> b"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    53
  by (auto simp: bequiv_up_to_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 bequiv_up_to_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    56
  "(P \<Turnstile> b <\<sim>> b') = (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_trans [trans]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    60
  "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
    61
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    62
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    63
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    64
lemma equiv_up_to_hoare:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    65
  "P' \<Turnstile> c \<sim> c' \<Longrightarrow> (\<And>s. P s \<Longrightarrow> P' s) \<Longrightarrow> (\<Turnstile> {P} c {Q}) = (\<Turnstile> {P} c' {Q})"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    66
  unfolding hoare_valid_def equiv_up_to_def by blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    67
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    68
lemma equiv_up_to_hoare_eq:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    69
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> (\<Turnstile> {P} c {Q}) = (\<Turnstile> {P} c' {Q})"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    70
  by (rule equiv_up_to_hoare)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    71
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    72
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    73
lemma equiv_up_to_semi:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    74
  "P \<Turnstile> c \<sim> c' \<Longrightarrow> Q \<Turnstile> d \<sim> d' \<Longrightarrow> \<Turnstile> {P} c {Q} \<Longrightarrow>
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    75
  P \<Turnstile> (c; d) \<sim> (c'; d')"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    76
  by (clarsimp simp: equiv_up_to_def hoare_valid_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    77
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    78
lemma equiv_up_to_while_lemma:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    79
  shows "(d,s) \<Rightarrow> s' \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    80
         P \<Turnstile> b <\<sim>> b' \<Longrightarrow>
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    81
         (\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c' \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    82
         \<Turnstile> {\<lambda>s. P s \<and> bval b s} c {P} \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    83
         P s \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    84
         d = WHILE b DO c \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
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)
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
    88
  note IH = WhileTrue.IH(2) [OF WhileTrue.prems(1-3) _ WhileTrue.prems(5)]
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
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    95
  have "(\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c'" by simp
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
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   100
  have "\<Turnstile> {\<lambda>s. P s \<and> bval b s} c {P}" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   101
  with `P s1` `bval b s1` `(c, s1) \<Rightarrow> s2`
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   102
  have "P s2" by (simp add: hoare_valid_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   103
  hence "(WHILE b' DO c', s2) \<Rightarrow> s3" by (rule IH)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   104
  ultimately 
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)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44261
diff changeset
   109
qed (fastforce simp: equiv_up_to_def bequiv_up_to_def hoare_valid_def)+
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   110
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   111
lemma bequiv_context_subst:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   112
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> (P s \<and> bval b s) = (P s \<and> bval b' s)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   113
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   114
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   115
lemma equiv_up_to_while:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   116
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> (\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c' \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   117
   \<Turnstile> {\<lambda>s. P s \<and> bval b s} c {P} \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   118
   P \<Turnstile> WHILE b DO c \<sim> WHILE b' DO c'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   119
  apply (safe intro!: equiv_up_toI)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   120
   apply (auto intro: equiv_up_to_while_lemma)[1]
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   121
  apply (simp add: equiv_up_to_hoare_eq bequiv_context_subst)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   122
  apply (drule equiv_up_to_sym [THEN iffD1])
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   123
  apply (drule bequiv_up_to_sym [THEN iffD1])
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   124
  apply (auto intro: equiv_up_to_while_lemma)[1]
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   125
  done
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   126
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   127
lemma equiv_up_to_while_weak:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   128
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> P \<Turnstile> c \<sim> c' \<Longrightarrow> \<Turnstile> {P} c {P} \<Longrightarrow> 
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   129
   P \<Turnstile> WHILE b DO c \<sim> WHILE b' DO c'"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44261
diff changeset
   130
  by (fastforce elim!: equiv_up_to_while equiv_up_to_weaken 
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   131
               simp: hoare_valid_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   132
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   133
lemma equiv_up_to_if:
44261
e44f465c00a1 HOL-IMP: respect set/pred distinction
huffman
parents: 44070
diff changeset
   134
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> (\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c' \<Longrightarrow> (\<lambda>s. P s \<and> \<not>bval b s) \<Turnstile> d \<sim> d' \<Longrightarrow>
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   135
   P \<Turnstile> IF b THEN c ELSE d \<sim> IF b' THEN c' ELSE d'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   136
  by (auto simp: bequiv_up_to_def equiv_up_to_def)
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 equiv_up_to_if_weak:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   139
  "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
   140
   P \<Turnstile> IF b THEN c ELSE d \<sim> IF b' THEN c' ELSE d'"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44261
diff changeset
   141
  by (fastforce elim!: equiv_up_to_if equiv_up_to_weaken)
44070
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_if_True [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   144
  "(\<And>s. P s \<Longrightarrow> bval b s) \<Longrightarrow> P \<Turnstile> IF b THEN c1 ELSE c2 \<sim> c1"
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
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   147
lemma equiv_up_to_if_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   148
  "(\<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
   149
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   150
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   151
lemma equiv_up_to_while_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   152
  "(\<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
   153
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   154
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   155
lemma while_never: "(c, s) \<Rightarrow> u \<Longrightarrow> c \<noteq> WHILE (B True) DO c'"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   156
 by (induct rule: big_step_induct) auto
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   157
  
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   158
lemma equiv_up_to_while_True [intro!,simp]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   159
  "P \<Turnstile> WHILE B True DO c \<sim> WHILE B True DO SKIP"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   160
  unfolding equiv_up_to_def
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   161
  by (blast dest: while_never)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   162
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   163
44261
e44f465c00a1 HOL-IMP: respect set/pred distinction
huffman
parents: 44070
diff changeset
   164
end