src/HOL/IMP/Expr.ML
author nipkow
Tue, 10 Sep 1996 20:10:29 +0200
changeset 1973 8c94c9a5be10
parent 1729 e4f8682eea2e
child 4089 96fba19bcbe2
permissions -rw-r--r--
Converted proofs to use default clasets.
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.ML
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
Arithmetic expressions and Boolean expressions.
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     7
Not used in the rest of the language, but included for completeness.
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     8
*)
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
     9
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    10
open Expr;
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    11
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    12
val evala_elim_cases = map (evala.mk_cases aexp.simps)
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    13
   ["(N(n),sigma) -a-> i", "(X(x),sigma) -a-> i",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    14
    "(Op1 f e,sigma) -a-> i", "(Op2 f a1 a2,sigma)  -a-> i"
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    15
   ];
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    16
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    17
val evalb_elim_cases = map (evalb.mk_cases bexp.simps)
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    18
   ["(true,sigma) -b-> x", "(false,sigma) -b-> x",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    19
    "(ROp f a0 a1,sigma) -b-> x", "(noti(b),sigma) -b-> x",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    20
    "(b0 andi b1,sigma) -b-> x", "(b0 ori b1,sigma) -b-> x"
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
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    23
val evalb_simps = map (fn s => prove_goal Expr.thy s
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    24
    (fn _ => [fast_tac (HOL_cs addSIs evalb.intrs addSEs evalb_elim_cases) 1]))
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    25
  ["((true,sigma) -b-> w) = (w=True)",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    26
   "((false,sigma) -b-> w) = (w=False)",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    27
   "((ROp f a0 a1,sigma) -b-> w) = \
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    28
\   (? m. (a0,sigma) -a-> m & (? n. (a1,sigma) -a-> n & w = f m n))",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    29
   "((noti(b),sigma) -b-> w) = (? x. (b,sigma) -b-> x & w = (~x))",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    30
   "((b0 andi b1,sigma) -b-> w) = \
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    31
\   (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x&y)))",
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    32
   "((b0 ori b1,sigma) -b-> w) = \
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    33
\   (? x. (b0,sigma) -b-> x & (? y. (b1,sigma) -b-> y & w = (x|y)))"];
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    34
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    35
goal Expr.thy "!n. ((a,s) -a-> n) = (n = A a s)";
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    36
by (aexp.induct_tac "a" 1);                               (* struct. ind. *)
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    37
by (ALLGOALS Simp_tac);                                   (* rewr. Den.   *)
1973
8c94c9a5be10 Converted proofs to use default clasets.
nipkow
parents: 1729
diff changeset
    38
by (TRYALL (fast_tac (!claset addSIs (evala.intrs@prems)
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    39
                             addSEs evala_elim_cases)));
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    40
qed_spec_mp "aexp_iff";
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    41
1729
e4f8682eea2e Removal of special syntax for -a-> and -b->
paulson
parents: 1697
diff changeset
    42
goal Expr.thy "!w. ((b,s) -b-> w) = (w = B b s)";
1697
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    43
by (bexp.induct_tac "b" 1);
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    44
by (ALLGOALS(asm_simp_tac (!simpset addcongs [conj_cong]
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    45
                                    addsimps (aexp_iff::evalb_simps))));
687f0710c22d Arithemtic and boolean expressions are now in a separate theory.
nipkow
parents:
diff changeset
    46
qed_spec_mp "bexp_iff";