src/HOL/IMPP/Com.thy
author berghofe
Wed, 11 Jul 2007 11:29:44 +0200
changeset 23756 14008ce7df96
parent 23746 a455e69c31cc
child 27362 a6dc1769fdda
permissions -rw-r--r--
Adapted to changes in Predicate theory.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     1
(*  Title:    HOL/IMPP/Com.thy
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     2
    ID:       $Id$
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     3
    Author:   David von Oheimb (based on a theory by Tobias Nipkow et al), TUM
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     4
    Copyright 1999 TUM
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     5
*)
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     6
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
     7
header {* Semantics of arithmetic and boolean expressions, Syntax of commands *}
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
     8
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
     9
theory Com
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    10
imports Main
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    11
begin
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    12
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    13
types    val = nat   (* for the meta theory, this may be anything, but with
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    14
                        current Isabelle, types cannot be refined later *)
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    15
typedecl glb
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    16
typedecl loc
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    17
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    18
consts
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    19
  Arg :: loc
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    20
  Res :: loc
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    21
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    22
datatype vname  = Glb glb | Loc loc
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    23
types    globs  = "glb => val"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    24
         locals = "loc => val"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    25
datatype state  = st globs locals
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    26
(* for the meta theory, the following would be sufficient:
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    27
typedecl state
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    28
consts   st :: "[globs , locals] => state"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    29
*)
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    30
types    aexp   = "state => val"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    31
         bexp   = "state => bool"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    32
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    33
typedecl pname
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    34
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    35
datatype com
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    36
      = SKIP
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    37
      | Ass   vname aexp        ("_:==_"                [65, 65    ] 60)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    38
      | Local loc aexp com      ("LOCAL _:=_ IN _"      [65,  0, 61] 60)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    39
      | Semi  com  com          ("_;; _"                [59, 60    ] 59)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    40
      | Cond  bexp com com      ("IF _ THEN _ ELSE _"   [65, 60, 61] 60)
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    41
      | While bexp com          ("WHILE _ DO _"         [65,     61] 60)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    42
      | BODY  pname
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    43
      | Call  vname pname aexp  ("_:=CALL _'(_')"       [65, 65,  0] 60)
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    44
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    45
consts bodies :: "(pname  *  com) list"(* finitely many procedure definitions *)
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    46
       body   :: " pname ~=> com"
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    47
defs   body_def: "body == map_of bodies"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    48
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    49
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    50
(* Well-typedness: all procedures called must exist *)
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    51
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    52
inductive WT  :: "com => bool" where
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    53
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    54
    Skip:    "WT SKIP"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    55
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    56
  | Assign:  "WT (X :== a)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    57
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    58
  | Local:   "WT c ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    59
              WT (LOCAL Y := a IN c)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    60
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    61
  | Semi:    "[| WT c0; WT c1 |] ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    62
              WT (c0;; c1)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    63
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    64
  | If:      "[| WT c0; WT c1 |] ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    65
              WT (IF b THEN c0 ELSE c1)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    66
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    67
  | While:   "WT c ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    68
              WT (WHILE b DO c)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    69
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    70
  | Body:    "body pn ~= None ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    71
              WT (BODY pn)"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    72
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 19803
diff changeset
    73
  | Call:    "WT (BODY pn) ==>
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    74
              WT (X:=CALL pn(a))"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    75
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    76
inductive_cases WTs_elim_cases:
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    77
  "WT SKIP"  "WT (X:==a)"  "WT (LOCAL Y:=a IN c)"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    78
  "WT (c1;;c2)"  "WT (IF b THEN c1 ELSE c2)"  "WT (WHILE b DO c)"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    79
  "WT (BODY P)"  "WT (X:=CALL P(a))"
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    80
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    81
constdefs
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
    82
  WT_bodies :: bool
17477
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    83
  "WT_bodies == !(pn,b):set bodies. WT b"
ceb42ea2f223 converted to Isar theory format;
wenzelm
parents: 12338
diff changeset
    84
19803
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    85
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    86
ML {* val make_imp_tac = EVERY'[rtac mp, fn i => atac (i+1), etac thin_rl] *}
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    87
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    88
lemma finite_dom_body: "finite (dom body)"
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    89
apply (unfold body_def)
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    90
apply (rule finite_dom_map_of)
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    91
done
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    92
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    93
lemma WT_bodiesD: "[| WT_bodies; body pn = Some b |] ==> WT b"
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    94
apply (unfold WT_bodies_def body_def)
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    95
apply (drule map_of_SomeD)
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    96
apply fast
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    97
done
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    98
aa2581752afb removed obsolete ML files;
wenzelm
parents: 17477
diff changeset
    99
declare WTs_elim_cases [elim!]
8177
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
   100
e59e93ad85eb added IMPP to HOL
oheimb
parents:
diff changeset
   101
end