src/HOL/Induct/ABexp.thy
author wenzelm
Fri, 25 Nov 2005 21:14:34 +0100
changeset 18260 5597cfcecd49
parent 16417 9bc16273c2d4
child 36862 952b2b102a0a
permissions -rw-r--r--
tuned induct proofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Induct/ABexp.thy
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     2
    ID:         $Id$
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     4
*)
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     5
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
     6
header {* Arithmetic and boolean expressions *}
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14981
diff changeset
     8
theory ABexp imports Main begin
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     9
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    10
datatype 'a aexp =
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    11
    IF "'a bexp"  "'a aexp"  "'a aexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    12
  | Sum "'a aexp"  "'a aexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    13
  | Diff "'a aexp"  "'a aexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    14
  | Var 'a
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    15
  | Num nat
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    16
and 'a bexp =
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    17
    Less "'a aexp"  "'a aexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    18
  | And "'a bexp"  "'a bexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    19
  | Neg "'a bexp"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    20
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    21
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    22
text {* \medskip Evaluation of arithmetic and boolean expressions *}
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    23
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    24
consts
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    25
  evala :: "('a => nat) => 'a aexp => nat"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    26
  evalb :: "('a => nat) => 'a bexp => bool"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    27
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    28
primrec
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    29
  "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"
5802
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    30
  "evala env (Sum a1 a2) = evala env a1 + evala env a2"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    31
  "evala env (Diff a1 a2) = evala env a1 - evala env a2"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    32
  "evala env (Var v) = env v"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    33
  "evala env (Num n) = n"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    34
5802
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    35
  "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    36
  "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    37
  "evalb env (Neg b) = (\<not> evalb env b)"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    38
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    39
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    40
text {* \medskip Substitution on arithmetic and boolean expressions *}
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    41
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    42
consts
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    43
  substa :: "('a => 'b aexp) => 'a aexp => 'b aexp"
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    44
  substb :: "('a => 'b aexp) => 'a bexp => 'b bexp"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    45
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    46
primrec
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    47
  "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"
5802
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    48
  "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    49
  "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    50
  "substa f (Var v) = f v"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    51
  "substa f (Num n) = Num n"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    52
5802
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    53
  "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    54
  "substb f (And b1 b2) = And (substb f b1) (substb f b2)"
614f2f30781a Shortened names and added new thm.
nipkow
parents: 5737
diff changeset
    55
  "substb f (Neg b) = Neg (substb f b)"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    56
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    57
lemma subst1_aexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    58
  "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    59
and subst1_bexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    60
  "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    61
    --  {* one variable *}
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    62
  by (induct a and b) simp_all
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    63
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    64
lemma subst_all_aexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    65
  "evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    66
and subst_all_bexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    67
  "evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    68
  by (induct a and b) auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    69
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    70
end