explicitly state equivalence relation for sim; tweak syntax of sem_equiv
authorkleing
Thu, 16 May 2013 13:49:18 +1000
changeset 52021 59963cda805a
parent 52020 db08e493cf44
child 52022 e0a83bd90bb1
child 52023 b477be38a7bb
explicitly state equivalence relation for sim; tweak syntax of sem_equiv
src/HOL/IMP/Big_Step.thy
src/HOL/IMP/Sem_Equiv.thy
--- a/src/HOL/IMP/Big_Step.thy	Thu May 16 02:13:42 2013 +0200
+++ b/src/HOL/IMP/Big_Step.thy	Thu May 16 13:49:18 2013 +1000
@@ -226,6 +226,13 @@
 lemma sim_while_cong: "c \<sim> c' \<Longrightarrow> WHILE b DO c \<sim> WHILE b DO c'"
 by (metis sim_while_cong_aux)
 
+text {* Command equivalence is an equivalence relation, i.e.\ it is
+reflexive, symmetric, and transitive. Because we used an abbreviation
+above, Isabelle derives this automatically. *}
+
+lemma sim_refl:  "c \<sim> c" by simp
+lemma sim_sym:   "(c \<sim> c') = (c' \<sim> c)" by auto
+lemma sim_trans: "c \<sim> c' \<Longrightarrow> c' \<sim> c'' \<Longrightarrow> c \<sim> c''" by auto
 
 subsection "Execution is deterministic"
 
--- a/src/HOL/IMP/Sem_Equiv.thy	Thu May 16 02:13:42 2013 +0200
+++ b/src/HOL/IMP/Sem_Equiv.thy	Thu May 16 13:49:18 2013 +1000
@@ -7,12 +7,12 @@
 type_synonym assn = "state \<Rightarrow> bool"
 
 definition
-  equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" ("_ \<Turnstile> _ \<sim> _" [60,0,10] 60)
+  equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" ("_ \<Turnstile> _ \<sim> _" [50,0,10] 50)
 where
   "P \<Turnstile> c \<sim> c' \<equiv> \<forall>s s'. P s \<longrightarrow> (c,s) \<Rightarrow> s' \<longleftrightarrow> (c',s) \<Rightarrow> s'"
 
 definition
-  bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" ("_ \<Turnstile> _ <\<sim>> _" [60,0,10] 60)
+  bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" ("_ \<Turnstile> _ <\<sim>> _" [50,0,10] 50)
 where
   "P \<Turnstile> b <\<sim>> b' \<equiv> \<forall>s. P s \<longrightarrow> bval b s = bval b' s"