src/HOLCF/IMP/Denotational.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 12600 30ec65eaaf5f
child 19737 3b8920131be2
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
     1
(*  Title:      HOLCF/IMP/Denotational.thy
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
12600
wenzelm
parents: 12431
diff changeset
     3
    Author:     Tobias Nipkow and Robert Sandner, TUM
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
     5
*)
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
     6
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
     7
header "Denotational Semantics of Commands in HOLCF"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12600
diff changeset
     9
theory Denotational imports HOLCF Natural begin
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    10
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    11
subsection "Definition"
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
    12
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
    13
constdefs
12600
wenzelm
parents: 12431
diff changeset
    14
  dlift :: "(('a::type) discr -> 'b::pcpo) => ('a lift -> 'b)"
wenzelm
parents: 12431
diff changeset
    15
  "dlift f == (LAM x. case x of UU => UU | Def y => f\<cdot>(Discr y))"
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
    16
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
    17
consts D :: "com => state discr -> state lift"
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
    18
5192
704dd3a6d47d Adapted to new datatype package.
berghofe
parents: 4904
diff changeset
    19
primrec
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    20
  "D(\<SKIP>) = (LAM s. Def(undiscr s))"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    21
  "D(X :== a) = (LAM s. Def((undiscr s)[X \<mapsto> a(undiscr s)]))"
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
    22
  "D(c0 ; c1) = (dlift(D c1) oo (D c0))"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    23
  "D(\<IF> b \<THEN> c1 \<ELSE> c2) =
12600
wenzelm
parents: 12431
diff changeset
    24
        (LAM s. if b (undiscr s) then (D c1)\<cdot>s else (D c2)\<cdot>s)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    25
  "D(\<WHILE> b \<DO> c) =
12600
wenzelm
parents: 12431
diff changeset
    26
        fix\<cdot>(LAM w s. if b (undiscr s) then (dlift w)\<cdot>((D c)\<cdot>s)
2841
c2508f4ab739 Added "discrete" CPOs and modified IMP to use those rather than "lift"
nipkow
parents: 2798
diff changeset
    27
                      else Def(undiscr s))"
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
    28
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    29
subsection
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    30
  "Equivalence of Denotational Semantics in HOLCF and Evaluation Semantics in HOL"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    31
12600
wenzelm
parents: 12431
diff changeset
    32
lemma dlift_Def [simp]: "dlift f\<cdot>(Def x) = f\<cdot>(Discr x)"
wenzelm
parents: 12431
diff changeset
    33
  by (simp add: dlift_def)
wenzelm
parents: 12431
diff changeset
    34
wenzelm
parents: 12431
diff changeset
    35
lemma cont_dlift [iff]: "cont (%f. dlift f)"
wenzelm
parents: 12431
diff changeset
    36
  by (simp add: dlift_def)
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    37
12600
wenzelm
parents: 12431
diff changeset
    38
lemma dlift_is_Def [simp]:
wenzelm
parents: 12431
diff changeset
    39
    "(dlift f\<cdot>l = Def y) = (\<exists>x. l = Def x \<and> f\<cdot>(Discr x) = Def y)"
wenzelm
parents: 12431
diff changeset
    40
  by (simp add: dlift_def split: lift.split)
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    41
12600
wenzelm
parents: 12431
diff changeset
    42
lemma eval_implies_D: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t ==> D c\<cdot>(Discr s) = (Def t)"
wenzelm
parents: 12431
diff changeset
    43
  apply (induct set: evalc)
wenzelm
parents: 12431
diff changeset
    44
        apply simp_all
wenzelm
parents: 12431
diff changeset
    45
   apply (subst fix_eq)
wenzelm
parents: 12431
diff changeset
    46
   apply simp
wenzelm
parents: 12431
diff changeset
    47
  apply (subst fix_eq)
wenzelm
parents: 12431
diff changeset
    48
  apply simp
wenzelm
parents: 12431
diff changeset
    49
  done
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    50
12600
wenzelm
parents: 12431
diff changeset
    51
lemma D_implies_eval: "!s t. D c\<cdot>(Discr s) = (Def t) --> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
wenzelm
parents: 12431
diff changeset
    52
  apply (induct c)
wenzelm
parents: 12431
diff changeset
    53
      apply simp
wenzelm
parents: 12431
diff changeset
    54
     apply simp
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    55
    apply force
12600
wenzelm
parents: 12431
diff changeset
    56
   apply (simp (no_asm))
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    57
   apply force
12600
wenzelm
parents: 12431
diff changeset
    58
  apply (simp (no_asm))
wenzelm
parents: 12431
diff changeset
    59
  apply (rule fix_ind)
wenzelm
parents: 12431
diff changeset
    60
    apply (fast intro!: adm_lemmas adm_chfindom ax_flat)
wenzelm
parents: 12431
diff changeset
    61
   apply (simp (no_asm))
wenzelm
parents: 12431
diff changeset
    62
  apply (simp (no_asm))
wenzelm
parents: 12431
diff changeset
    63
  apply safe
wenzelm
parents: 12431
diff changeset
    64
  apply fast
wenzelm
parents: 12431
diff changeset
    65
  done
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    66
12600
wenzelm
parents: 12431
diff changeset
    67
theorem D_is_eval: "(D c\<cdot>(Discr s) = (Def t)) = (\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t)"
wenzelm
parents: 12431
diff changeset
    68
  by (fast elim!: D_implies_eval [rule_format] eval_implies_D)
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    69
2798
f84be65745b2 The HOLCF-based den. sem. of IMP.
nipkow
parents:
diff changeset
    70
end