| author | krauss | 
| Sun, 23 May 2010 22:56:45 +0200 | |
| changeset 37095 | 805d18dae026 | 
| parent 34055 | fdf294ee08b2 | 
| child 41589 | bbd861837ebc | 
| permissions | -rw-r--r-- | 
| 1476 | 1  | 
(* Title: HOL/IMP/Denotation.thy  | 
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 1476 | 3  | 
Author: Heiko Loetzbeyer & Robert Sandner, TUM  | 
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
4  | 
Copyright 1994 TUM  | 
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
6  | 
|
| 12431 | 7  | 
header "Denotational Semantics of Commands"  | 
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
8  | 
|
| 16417 | 9  | 
theory Denotation imports Natural begin  | 
| 12431 | 10  | 
|
11  | 
types com_den = "(state\<times>state)set"  | 
|
| 1696 | 12  | 
|
| 27362 | 13  | 
definition  | 
14  | 
Gamma :: "[bexp,com_den] => (com_den => com_den)" where  | 
|
| 
32235
 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 
krauss 
parents: 
27362 
diff
changeset
 | 
15  | 
  "Gamma b cd = (\<lambda>phi. {(s,t). (s,t) \<in> (cd O phi) \<and> b s} \<union>
 | 
| 12434 | 16  | 
                       {(s,t). s=t \<and> \<not>b s})"
 | 
| 18372 | 17  | 
|
| 27362 | 18  | 
primrec C :: "com => com_den"  | 
19  | 
where  | 
|
| 12431 | 20  | 
C_skip: "C \<SKIP> = Id"  | 
| 27362 | 21  | 
| C_assign: "C (x :== a) = {(s,t). t = s[x\<mapsto>a(s)]}"
 | 
| 
32235
 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 
krauss 
parents: 
27362 
diff
changeset
 | 
22  | 
| C_comp: "C (c0;c1) = C(c0) O C(c1)"  | 
| 27362 | 23  | 
| C_if:     "C (\<IF> b \<THEN> c1 \<ELSE> c2) = {(s,t). (s,t) \<in> C c1 \<and> b s} \<union>
 | 
| 12434 | 24  | 
                                                {(s,t). (s,t) \<in> C c2 \<and> \<not>b s}"
 | 
| 27362 | 25  | 
| C_while: "C(\<WHILE> b \<DO> c) = lfp (Gamma b (C c))"  | 
| 12431 | 26  | 
|
27  | 
||
28  | 
(**** mono (Gamma(b,c)) ****)  | 
|
29  | 
||
30  | 
lemma Gamma_mono: "mono (Gamma b c)"  | 
|
31  | 
by (unfold Gamma_def mono_def) fast  | 
|
32  | 
||
33  | 
lemma C_While_If: "C(\<WHILE> b \<DO> c) = C(\<IF> b \<THEN> c;\<WHILE> b \<DO> c \<ELSE> \<SKIP>)"  | 
|
| 18372 | 34  | 
apply simp  | 
| 15481 | 35  | 
apply (subst lfp_unfold [OF Gamma_mono])  --{*lhs only*}
 | 
36  | 
apply (simp add: Gamma_def)  | 
|
| 12431 | 37  | 
done  | 
38  | 
||
39  | 
(* Operational Semantics implies Denotational Semantics *)  | 
|
40  | 
||
41  | 
lemma com1: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t \<Longrightarrow> (s,t) \<in> C(c)"  | 
|
42  | 
(* start with rule induction *)  | 
|
| 18372 | 43  | 
apply (induct set: evalc)  | 
| 12431 | 44  | 
apply auto  | 
45  | 
(* while *)  | 
|
46  | 
apply (unfold Gamma_def)  | 
|
47  | 
apply (subst lfp_unfold[OF Gamma_mono, simplified Gamma_def])  | 
|
48  | 
apply fast  | 
|
49  | 
apply (subst lfp_unfold[OF Gamma_mono, simplified Gamma_def])  | 
|
| 34055 | 50  | 
apply auto  | 
| 12431 | 51  | 
done  | 
52  | 
||
53  | 
(* Denotational Semantics implies Operational Semantics *)  | 
|
54  | 
||
| 18372 | 55  | 
lemma com2: "(s,t) \<in> C(c) \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 20503 | 56  | 
apply (induct c arbitrary: s t)  | 
| 34055 | 57  | 
apply auto  | 
58  | 
apply blast  | 
|
| 12431 | 59  | 
|
60  | 
(* while *)  | 
|
| 23746 | 61  | 
apply (erule lfp_induct2 [OF _ Gamma_mono])  | 
| 12431 | 62  | 
apply (unfold Gamma_def)  | 
| 34055 | 63  | 
apply auto  | 
| 12431 | 64  | 
done  | 
65  | 
||
66  | 
||
67  | 
(**** Proof of Equivalence ****)  | 
|
68  | 
||
69  | 
lemma denotational_is_natural: "(s,t) \<in> C(c) = (\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t)"  | 
|
| 18372 | 70  | 
by (fast elim: com2 dest: com1)  | 
| 
924
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
806721cfbf46
new version of HOL/IMP with curried function application
 
clasohm 
parents:  
diff
changeset
 | 
72  | 
end  |