author | paulson |
Tue, 01 Feb 2005 18:01:57 +0100 | |
changeset 15481 | fc075ae929e4 |
parent 12434 | ff2efde4574d |
child 16417 | 9bc16273c2d4 |
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 |
|
12431 | 9 |
theory Denotation = Natural: |
10 |
||
11 |
types com_den = "(state\<times>state)set" |
|
1696 | 12 |
|
13 |
constdefs |
|
12431 | 14 |
Gamma :: "[bexp,com_den] => (com_den => com_den)" |
15 |
"Gamma b cd == (\<lambda>phi. {(s,t). (s,t) \<in> (phi O cd) \<and> b s} \<union> |
|
12434 | 16 |
{(s,t). s=t \<and> \<not>b s})" |
1696 | 17 |
|
18 |
consts |
|
12431 | 19 |
C :: "com => com_den" |
924
806721cfbf46
new version of HOL/IMP with curried function application
clasohm
parents:
diff
changeset
|
20 |
|
5183 | 21 |
primrec |
12431 | 22 |
C_skip: "C \<SKIP> = Id" |
23 |
C_assign: "C (x :== a) = {(s,t). t = s[x\<mapsto>a(s)]}" |
|
24 |
C_comp: "C (c0;c1) = C(c1) O C(c0)" |
|
25 |
C_if: "C (\<IF> b \<THEN> c1 \<ELSE> c2) = {(s,t). (s,t) \<in> C c1 \<and> b s} \<union> |
|
12434 | 26 |
{(s,t). (s,t) \<in> C c2 \<and> \<not>b s}" |
12431 | 27 |
C_while: "C(\<WHILE> b \<DO> c) = lfp (Gamma b (C c))" |
28 |
||
29 |
||
30 |
(**** mono (Gamma(b,c)) ****) |
|
31 |
||
32 |
lemma Gamma_mono: "mono (Gamma b c)" |
|
33 |
by (unfold Gamma_def mono_def) fast |
|
34 |
||
35 |
lemma C_While_If: "C(\<WHILE> b \<DO> c) = C(\<IF> b \<THEN> c;\<WHILE> b \<DO> c \<ELSE> \<SKIP>)" |
|
12434 | 36 |
apply (simp (no_asm)) |
15481 | 37 |
apply (subst lfp_unfold [OF Gamma_mono]) --{*lhs only*} |
38 |
apply (simp add: Gamma_def) |
|
12431 | 39 |
done |
40 |
||
41 |
(* Operational Semantics implies Denotational Semantics *) |
|
42 |
||
43 |
lemma com1: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t \<Longrightarrow> (s,t) \<in> C(c)" |
|
44 |
(* start with rule induction *) |
|
45 |
apply (erule evalc.induct) |
|
46 |
apply auto |
|
47 |
(* while *) |
|
48 |
apply (unfold Gamma_def) |
|
49 |
apply (subst lfp_unfold[OF Gamma_mono, simplified Gamma_def]) |
|
50 |
apply fast |
|
51 |
apply (subst lfp_unfold[OF Gamma_mono, simplified Gamma_def]) |
|
52 |
apply fast |
|
53 |
done |
|
54 |
||
55 |
(* Denotational Semantics implies Operational Semantics *) |
|
56 |
||
57 |
lemma com2 [rule_format]: "\<forall>s t. (s,t)\<in>C(c) \<longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" |
|
58 |
apply (induct_tac "c") |
|
59 |
||
60 |
apply (simp_all (no_asm_use)) |
|
61 |
apply fast |
|
62 |
apply fast |
|
63 |
||
64 |
(* while *) |
|
65 |
apply (intro strip) |
|
66 |
apply (erule lfp_induct [OF _ Gamma_mono]) |
|
67 |
apply (unfold Gamma_def) |
|
68 |
apply fast |
|
69 |
done |
|
70 |
||
71 |
||
72 |
(**** Proof of Equivalence ****) |
|
73 |
||
74 |
lemma denotational_is_natural: "(s,t) \<in> C(c) = (\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t)" |
|
75 |
apply (fast elim: com2 dest: com1) |
|
76 |
done |
|
924
806721cfbf46
new version of HOL/IMP with curried function application
clasohm
parents:
diff
changeset
|
77 |
|
806721cfbf46
new version of HOL/IMP with curried function application
clasohm
parents:
diff
changeset
|
78 |
end |