src/HOL/IMP/Expr.thy
author berghofe
Mon, 26 Oct 2009 14:53:33 +0100
changeset 33189 82a40677c1f8
parent 27362 a6dc1769fdda
child 37736 2bf3a2cb5e58
permissions -rw-r--r--
Added Pattern.thy to Nominal/Examples.
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
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12546
diff changeset
     9
theory Expr imports Main begin
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"
12546
wenzelm
parents: 12431
diff changeset
    17
typedecl loc
12431
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"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    21
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    22
datatype
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    23
  aexp = N nat
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    24
       | X loc
12546
wenzelm
parents: 12431
diff changeset
    25
       | Op1 "nat => nat" aexp
wenzelm
parents: 12431
diff changeset
    26
       | Op2 "nat => nat => nat" aexp aexp
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    27
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    28
subsection "Evaluation of arithmetic expressions"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    29
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    30
inductive
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    31
  evala :: "[aexp*state,nat] => bool"  (infixl "-a->" 50)
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    32
where
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    33
  N:   "(N(n),s) -a-> n"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    34
| X:   "(X(x),s) -a-> s(x)"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    35
| Op1: "(e,s) -a-> n ==> (Op1 f e,s) -a-> f(n)"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    36
| Op2: "[| (e0,s) -a-> n0;  (e1,s)  -a-> n1 |]
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    37
        ==> (Op2 f e0 e1,s) -a-> f n0 n1"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    38
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    39
lemmas [intro] = N X Op1 Op2
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    40
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    41
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    42
subsection "Boolean expressions"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    43
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    44
datatype
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    45
  bexp = true
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    46
       | false
12546
wenzelm
parents: 12431
diff changeset
    47
       | ROp  "nat => nat => bool" aexp aexp
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    48
       | noti bexp
22818
c0695a818c09 eliminated unnamed infixes;
wenzelm
parents: 20503
diff changeset
    49
       | andi bexp bexp         (infixl "andi" 60)
c0695a818c09 eliminated unnamed infixes;
wenzelm
parents: 20503
diff changeset
    50
       | ori  bexp bexp         (infixl "ori" 60)
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    51
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    52
subsection "Evaluation of boolean expressions"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    53
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    54
inductive
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    55
  evalb :: "[bexp*state,bool] => bool"  (infixl "-b->" 50)
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    56
  -- "avoid clash with ML constructors true, false"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    57
where
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    58
  tru:   "(true,s) -b-> True"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    59
| fls:   "(false,s) -b-> False"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    60
| ROp:   "[| (a0,s) -a-> n0; (a1,s) -a-> n1 |]
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    61
          ==> (ROp f a0 a1,s) -b-> f n0 n1"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    62
| noti:  "(b,s) -b-> w ==> (noti(b),s) -b-> (~w)"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    63
| andi:  "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    64
          ==> (b0 andi b1,s) -b-> (w0 & w1)"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
    65
| ori:   "[| (b0,s) -b-> w0; (b1,s) -b-> w1 |]
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    66
          ==> (b0 ori b1,s) -b-> (w0 | w1)"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    67
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    68
lemmas [intro] = tru fls ROp noti andi ori
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    69
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    70
subsection "Denotational semantics of arithmetic and boolean expressions"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    71
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    72
primrec A :: "aexp => state => nat"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    73
where
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    74
  "A(N(n)) = (%s. n)"
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    75
| "A(X(x)) = (%s. s(x))"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    76
| "A(Op1 f a) = (%s. f(A a s))"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    77
| "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
    78
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    79
primrec B :: "bexp => state => bool"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    80
where
1900
c7a869229091 Simplified primrec definitions.
berghofe
parents: 1789
diff changeset
    81
  "B(true) = (%s. True)"
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    82
| "B(false) = (%s. False)"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    83
| "B(ROp f a0 a1) = (%s. f (A a0 s) (A a1 s))"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    84
| "B(noti(b)) = (%s. ~(B b s))"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    85
| "B(b0 andi b1) = (%s. (B b0 s) & (B b1 s))"
a6dc1769fdda modernized specifications;
wenzelm
parents: 23746
diff changeset
    86
| "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
    87
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    88
lemma [simp]: "(N(n),s) -a-> n' = (n = n')"
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    89
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    90
12546
wenzelm
parents: 12431
diff changeset
    91
lemma [simp]: "(X(x),sigma) -a-> i = (i = sigma x)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    92
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    93
12546
wenzelm
parents: 12431
diff changeset
    94
lemma   [simp]:
wenzelm
parents: 12431
diff changeset
    95
  "(Op1 f e,sigma) -a-> i = (\<exists>n. i = f n \<and> (e,sigma) -a-> n)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    96
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    97
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
    98
lemma [simp]:
12546
wenzelm
parents: 12431
diff changeset
    99
  "(Op2 f a1 a2,sigma) -a-> i =
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   100
  (\<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
   101
  by (rule,cases set: evala) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   102
12546
wenzelm
parents: 12431
diff changeset
   103
lemma [simp]: "((true,sigma) -b-> w) = (w=True)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   104
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   105
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   106
lemma [simp]:
12546
wenzelm
parents: 12431
diff changeset
   107
  "((false,sigma) -b-> w) = (w=False)"
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   108
  by (rule,cases set: evalb) auto
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   109
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   110
lemma [simp]:
12546
wenzelm
parents: 12431
diff changeset
   111
  "((ROp f a0 a1,sigma) -b-> w) =
wenzelm
parents: 12431
diff changeset
   112
  (? m. (a0,sigma) -a-> m & (? n. (a1,sigma) -a-> n & w = f m n))"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
   113
  by (rule,cases set: evalb) blast+
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   114
12546
wenzelm
parents: 12431
diff changeset
   115
lemma [simp]:
wenzelm
parents: 12431
diff changeset
   116
  "((noti(b),sigma) -b-> w) = (? x. (b,sigma) -b-> x & w = (~x))"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
   117
  by (rule,cases set: evalb) blast+
12546
wenzelm
parents: 12431
diff changeset
   118
wenzelm
parents: 12431
diff changeset
   119
lemma [simp]:
wenzelm
parents: 12431
diff changeset
   120
  "((b0 andi b1,sigma) -b-> w) =
wenzelm
parents: 12431
diff changeset
   121
  (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x&y)))"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
   122
  by (rule,cases set: evalb) blast+
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   123
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   124
lemma [simp]:
12546
wenzelm
parents: 12431
diff changeset
   125
  "((b0 ori b1,sigma) -b-> w) =
wenzelm
parents: 12431
diff changeset
   126
  (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 22818
diff changeset
   127
  by (rule,cases set: evalb) blast+
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   128
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   129
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   130
lemma aexp_iff: "((a,s) -a-> n) = (A a s = n)"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18372
diff changeset
   131
  by (induct a arbitrary: n) auto
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   132
12546
wenzelm
parents: 12431
diff changeset
   133
lemma bexp_iff:
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   134
  "((b,s) -b-> w) = (B b s = w)"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18372
diff changeset
   135
  by (induct b arbitrary: w) (auto simp add: aexp_iff)
12431
07ec657249e5 converted to Isar
kleing
parents: 12338
diff changeset
   136
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
   137
end