src/HOL/IMP/Sem_Equiv.thy
author haftmann
Fri, 01 Nov 2013 18:51:14 +0100
changeset 54230 b1d955791529
parent 52121 5b889b1b465b
child 54296 111ecbaa09f7
permissions -rw-r--r--
more simplification rules on unary and binary minus
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
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
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
     7
type_synonym assn = "state \<Rightarrow> bool"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
     8
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
     9
definition
52021
59963cda805a explicitly state equivalence relation for sim; tweak syntax of sem_equiv
kleing
parents: 48909
diff changeset
    10
  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
    11
where
52121
5b889b1b465b prefer object equality
kleing
parents: 52046
diff changeset
    12
  "(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
    13
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    14
definition
52021
59963cda805a explicitly state equivalence relation for sim; tweak syntax of sem_equiv
kleing
parents: 48909
diff changeset
    15
  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
    16
where
52121
5b889b1b465b prefer object equality
kleing
parents: 52046
diff changeset
    17
  "(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
    18
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    19
lemma equiv_up_to_True:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    20
  "((\<lambda>_. True) \<Turnstile> c \<sim> c') = (c \<sim> c')"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    21
  by (simp add: equiv_def equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    22
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    23
lemma equiv_up_to_weaken:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    24
  "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
    25
  by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    26
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    27
lemma equiv_up_toI:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    28
  "(\<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
    29
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    30
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    31
lemma equiv_up_toD1:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    32
  "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
    33
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    34
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    35
lemma equiv_up_toD2:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    36
  "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
    37
  by (unfold equiv_up_to_def) blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    38
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    39
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    40
lemma equiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    41
  "P \<Turnstile> c \<sim> c"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    42
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    43
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    44
lemma equiv_up_to_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    45
  "(P \<Turnstile> c \<sim> c') = (P \<Turnstile> c' \<sim> c)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    46
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    47
45218
f115540543d8 removed [trans] concept from basic material
kleing
parents: 45200
diff changeset
    48
lemma equiv_up_to_trans:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    49
  "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
    50
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    51
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 bequiv_up_to_refl [simp, intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    54
  "P \<Turnstile> b <\<sim>> b"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    55
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    56
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    57
lemma bequiv_up_to_sym:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    58
  "(P \<Turnstile> b <\<sim>> b') = (P \<Turnstile> b' <\<sim>> b)"
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    59
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    60
45218
f115540543d8 removed [trans] concept from basic material
kleing
parents: 45200
diff changeset
    61
lemma bequiv_up_to_trans:
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    62
  "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
    63
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    64
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    65
lemma bequiv_up_to_subst:
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    66
  "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
    67
  by (simp add: bequiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    68
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    69
47818
151d137f1095 renamed Semi to Seq
nipkow
parents: 45218
diff changeset
    70
lemma equiv_up_to_seq:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    71
  "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
    72
  (\<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
    73
  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
    74
  by (clarsimp simp: equiv_up_to_def) blast
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    75
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    76
lemma equiv_up_to_while_lemma:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    77
  shows "(d,s) \<Rightarrow> s' \<Longrightarrow>
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    78
         P \<Turnstile> b <\<sim>> b' \<Longrightarrow>
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    79
         (\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c' \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    80
         (\<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
    81
         P s \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    82
         d = WHILE b DO c \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    83
         (WHILE b' DO c', s) \<Rightarrow> s'"
45015
fdac1e9880eb Updated IMP to use new induction method
nipkow
parents: 44890
diff changeset
    84
proof (induction rule: big_step_induct)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    85
  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
    86
  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
    87
  from WhileTrue.prems
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    88
  have "P \<Turnstile> b <\<sim>> b'" by simp
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    89
  with `bval b s1` `P s1`
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    90
  have "bval b' s1" by (simp add: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    91
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    92
  from WhileTrue.prems
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    93
  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
    94
  with `bval b s1` `P s1` `(c, s1) \<Rightarrow> s2`
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    95
  have "(c', s1) \<Rightarrow> s2" by (simp add: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    96
  moreover
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
    97
  from WhileTrue.prems
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
    98
  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
    99
  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
   100
  have "P s2" by simp
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   101
  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
   102
  ultimately
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   103
  show ?case by blast
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   104
next
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   105
  case WhileFalse
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   106
  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
   107
qed (fastforce simp: equiv_up_to_def bequiv_up_to_def)+
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   108
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   109
lemma bequiv_context_subst:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   110
  "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
   111
  by (auto simp: bequiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   112
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   113
lemma equiv_up_to_while:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   114
  assumes b: "P \<Turnstile> b <\<sim>> b'"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   115
  assumes c: "(\<lambda>s. P s \<and> bval b s) \<Turnstile> c \<sim> c'"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   116
  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
   117
  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
   118
proof -
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   119
  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
   120
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   121
  from c b have c': "(\<lambda>s. P s \<and> bval b' s) \<Turnstile> c' \<sim> c"
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   122
    by (simp add: equiv_up_to_sym bequiv_context_subst)
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   123
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   124
  from I
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   125
  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
   126
    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
   127
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   128
  note equiv_up_to_while_lemma [OF _ b c]
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   129
       equiv_up_to_while_lemma [OF _ b' c']
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   130
  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
   131
qed
44070
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_while_weak:
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   134
  "P \<Turnstile> b <\<sim>> b' \<Longrightarrow> P \<Turnstile> c \<sim> c' \<Longrightarrow>
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   135
   (\<And>s s'. (c, s) \<Rightarrow> s' \<Longrightarrow> P s \<Longrightarrow> bval b s \<Longrightarrow> P s') \<Longrightarrow>
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   136
   P \<Turnstile> WHILE b DO c \<sim> WHILE b' DO c'"
48909
b2dea007e55e remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents: 47818
diff changeset
   137
  by (fastforce elim!: equiv_up_to_while equiv_up_to_weaken)
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:
44261
e44f465c00a1 HOL-IMP: respect set/pred distinction
huffman
parents: 44070
diff changeset
   140
  "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
   141
   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
   142
  by (auto simp: bequiv_up_to_def equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   143
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   144
lemma equiv_up_to_if_weak:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   145
  "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
   146
   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
   147
  by (fastforce elim!: equiv_up_to_if equiv_up_to_weaken)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   148
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   149
lemma equiv_up_to_if_True [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   150
  "(\<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
   151
  by (auto simp: equiv_up_to_def)
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   152
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   153
lemma equiv_up_to_if_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   154
  "(\<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
   155
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   156
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   157
lemma equiv_up_to_while_False [intro!]:
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   158
  "(\<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
   159
  by (auto simp: equiv_up_to_def)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   160
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
   161
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
   162
 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
   163
44070
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   164
lemma equiv_up_to_while_True [intro!,simp]:
45200
1f1897ac7877 renamed B to Bc
nipkow
parents: 45015
diff changeset
   165
  "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
   166
  unfolding equiv_up_to_def
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   167
  by (blast dest: while_never)
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   168
cebb7abb54b1 import constant folding theory into IMP
kleing
parents:
diff changeset
   169
44261
e44f465c00a1 HOL-IMP: respect set/pred distinction
huffman
parents: 44070
diff changeset
   170
end