src/HOL/IMP/Expr.thy
author kleing
Sun, 09 Dec 2001 14:35:36 +0100
changeset 12431 07ec657249e5
parent 12338 de0f4a63baa5
child 12546 0c90e581379f
permissions -rw-r--r--
converted to Isar
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/IMP/Expr.thy
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     3
    Author:     Heiko Loetzbeyer & Robert Sandner & Tobias Nipkow, TUM
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     4
    Copyright   1994 TUM
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     5
*)
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     6
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
     7
header "Expressions"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
     8
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
     9
theory Expr = Datatype:
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    10
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    11
text {*
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    12
  Arithmetic expressions and Boolean expressions.
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    13
  Not used in the rest of the language, but included for completeness.
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    14
*}
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    15
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    16
subsection "Arithmetic expressions"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    17
typedecl loc 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    18
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    19
types
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    20
  state = "loc => nat"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    21
  n2n = "nat => nat"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    22
  n2n2n = "nat => nat => nat"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    23
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    24
datatype
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    25
  aexp = N nat
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    26
       | X loc
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    27
       | Op1 n2n aexp
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    28
       | Op2 n2n2n aexp aexp
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    29
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    30
subsection "Evaluation of arithmetic expressions"
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    31
consts  evala    :: "((aexp*state) * nat) set"
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    32
       "-a->"    :: "[aexp*state,nat] => bool"         (infixl 50)
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    33
translations
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    34
    "aesig -a-> n" == "(aesig,n) : evala"
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1729
diff changeset
    35
inductive evala
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    36
  intros 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    37
  N:   "(N(n),s) -a-> n"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    38
  X:   "(X(x),s) -a-> s(x)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    39
  Op1: "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    40
  Op2: "[| (e0,s) -a-> n0;  (e1,s)  -a-> n1 |] 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    41
        ==> (Op2 f e0 e1,s) -a-> f n0 n1"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    42
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    43
lemmas [intro] = N X Op1 Op2
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    44
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    45
types n2n2b = "[nat,nat] => bool"
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    46
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    47
subsection "Boolean expressions"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    48
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    49
datatype
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    50
  bexp = true
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    51
       | false
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    52
       | ROp  n2n2b aexp aexp
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    53
       | noti bexp
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    54
       | andi bexp bexp         (infixl 60)
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    55
       | ori  bexp bexp         (infixl 60)
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    56
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    57
subsection "Evaluation of boolean expressions"
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    58
consts evalb    :: "((bexp*state) * bool)set"       
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    59
       "-b->"   :: "[bexp*state,bool] => bool"         (infixl 50)
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    60
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    61
translations
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    62
    "besig -b-> b" == "(besig,b) : evalb"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    63
1789
aade046ec6d5 Quotes now optional around inductive set
paulson
parents: 1729
diff changeset
    64
inductive evalb
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    65
  -- "avoid clash with ML constructors true, false"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    66
  intros
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    67
  tru:   "(true,s) -b-> True"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    68
  fls:   "(false,s) -b-> False"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    69
  ROp:   "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |] 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    70
          ==> (ROp f a0 a1,s) -b-> f n0 n1"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    71
  noti:  "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    72
  andi:  "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] 
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    73
          ==> (b0 andi b1,s) -b-> (w0 & w1)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    74
  ori:   "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |] 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    75
          ==> (b0 ori b1,s) -b-> (w0 | w1)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    76
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    77
lemmas [intro] = tru fls ROp noti andi ori
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    78
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    79
subsection "Denotational semantics of arithmetic and boolean expressions"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    80
consts
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    81
  A     :: "aexp => state => nat"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    82
  B     :: "bexp => state => bool"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    83
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5102
diff changeset
    84
primrec
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    85
  "A(N(n)) = (%s. n)"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    86
  "A(X(x)) = (%s. s(x))"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    87
  "A(Op1 f a) = (%s. f(A a s))"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    88
  "A(Op2 f a0 a1) = (%s. f (A a0 s) (A a1 s))"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    89
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 5102
diff changeset
    90
primrec
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    91
  "B(true) = (%s. True)"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    92
  "B(false) = (%s. False)"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    93
  "B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    94
  "B(noti(b)) = (%s. ~(B b s))"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    95
  "B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))"
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    96
  "B(b0 ori b1) = (%s. (B b0 s) | (B b1 s))"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    97
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    98
lemma [simp]: "(N(n),s) -a-> n' = (n = n')"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    99
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   100
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   101
lemma [simp]: "(X(x),sigma) -a-> i = (i = sigma x)" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   102
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   103
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   104
lemma	[simp]: 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   105
  "(Op1 f e,sigma) -a-> i = (\<exists>n. i = f n \<and> (e,sigma) -a-> n)" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   106
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   107
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   108
lemma [simp]:
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   109
  "(Op2 f a1 a2,sigma) -a-> i = 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   110
  (\<exists>n0 n1. i = f n0 n1 \<and> (a1, sigma) -a-> n0 \<and> (a2, sigma) -a-> n1)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   111
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   112
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   113
lemma [simp]: "((true,sigma) -b-> w) = (w=True)" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   114
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   115
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   116
lemma [simp]:
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   117
  "((false,sigma) -b-> w) = (w=False)" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   118
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   119
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   120
lemma [simp]:
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   121
  "((ROp f a0 a1,sigma) -b-> w) =   
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   122
  (? m. (a0,sigma) -a-> m & (? n. (a1,sigma) -a-> n & w = f m n))" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   123
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   124
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   125
lemma [simp]:  
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   126
  "((noti(b),sigma) -b-> w) = (? x. (b,sigma) -b-> x & w = (~x))" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   127
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   128
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   129
lemma [simp]:
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   130
  "((b0 andi b1,sigma) -b-> w) =   
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   131
  (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x&y)))" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   132
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   133
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   134
lemma [simp]:  
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   135
  "((b0 ori b1,sigma) -b-> w) =   
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   136
  (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))" 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   137
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   138
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   139
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   140
lemma aexp_iff [rule_format (no_asm)]: 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   141
  "!n. ((a,s) -a-> n) = (A a s = n)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   142
  apply (induct_tac "a")
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   143
  apply auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   144
  done
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   145
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   146
lemma bexp_iff [rule_format (no_asm)]: 
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   147
  "!w. ((b,s) -b-> w) = (B b s = w)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   148
  apply (induct_tac "b")
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   149
  apply (auto simp add: aexp_iff)
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   150
  done
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   151
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
   152
end
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
   153