| author | wenzelm | 
| Mon, 07 Jan 2002 18:58:35 +0100 | |
| changeset 12655 | b8c130dc46be | 
| parent 12171 | dc87f33db447 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 5737 | 1  | 
(* Title: HOL/Induct/ABexp.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Stefan Berghofer, TU Muenchen  | 
|
| 11649 | 4  | 
License: GPL (GNU GENERAL PUBLIC LICENSE)  | 
| 5737 | 5  | 
*)  | 
6  | 
||
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
7  | 
header {* Arithmetic and boolean expressions *}
 | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
8  | 
|
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
9  | 
theory ABexp = Main:  | 
| 5737 | 10  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
11  | 
datatype 'a aexp =  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
12  | 
IF "'a bexp" "'a aexp" "'a aexp"  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
13  | 
| Sum "'a aexp" "'a aexp"  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
14  | 
| Diff "'a aexp" "'a aexp"  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
15  | 
| Var 'a  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
16  | 
| Num nat  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
17  | 
and 'a bexp =  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
18  | 
Less "'a aexp" "'a aexp"  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
19  | 
| And "'a bexp" "'a bexp"  | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
20  | 
| Neg "'a bexp"  | 
| 5737 | 21  | 
|
22  | 
||
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
23  | 
text {* \medskip Evaluation of arithmetic and boolean expressions *}
 | 
| 5737 | 24  | 
|
25  | 
consts  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
26  | 
  evala :: "('a => nat) => 'a aexp => nat"
 | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
27  | 
  evalb :: "('a => nat) => 'a bexp => bool"
 | 
| 5737 | 28  | 
|
29  | 
primrec  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
30  | 
"evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"  | 
| 5802 | 31  | 
"evala env (Sum a1 a2) = evala env a1 + evala env a2"  | 
32  | 
"evala env (Diff a1 a2) = evala env a1 - evala env a2"  | 
|
33  | 
"evala env (Var v) = env v"  | 
|
34  | 
"evala env (Num n) = n"  | 
|
| 5737 | 35  | 
|
| 5802 | 36  | 
"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
 | 
37  | 
"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
 | 
38  | 
"evalb env (Neg b) = (\<not> evalb env b)"  | 
| 5737 | 39  | 
|
40  | 
||
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
41  | 
text {* \medskip Substitution on arithmetic and boolean expressions *}
 | 
| 5737 | 42  | 
|
43  | 
consts  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
44  | 
  substa :: "('a => 'b aexp) => 'a aexp => 'b aexp"
 | 
| 
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
45  | 
  substb :: "('a => 'b aexp) => 'a bexp => 'b bexp"
 | 
| 5737 | 46  | 
|
47  | 
primrec  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
48  | 
"substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"  | 
| 5802 | 49  | 
"substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"  | 
50  | 
"substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"  | 
|
51  | 
"substa f (Var v) = f v"  | 
|
52  | 
"substa f (Num n) = Num n"  | 
|
| 5737 | 53  | 
|
| 5802 | 54  | 
"substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"  | 
55  | 
"substb f (And b1 b2) = And (substb f b1) (substb f b2)"  | 
|
56  | 
"substb f (Neg b) = Neg (substb f b)"  | 
|
| 5737 | 57  | 
|
| 12171 | 58  | 
lemma subst1_aexp:  | 
59  | 
"evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"  | 
|
60  | 
and subst1_bexp:  | 
|
61  | 
"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
 | 
62  | 
    --  {* one variable *}
 | 
| 12171 | 63  | 
by (induct a and b) simp_all  | 
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
64  | 
|
| 12171 | 65  | 
lemma subst_all_aexp:  | 
66  | 
"evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"  | 
|
67  | 
and subst_all_bexp:  | 
|
68  | 
"evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"  | 
|
69  | 
by (induct a and b) auto  | 
|
| 
11046
 
b5f5942781a0
Induct: converted some theories to new-style format;
 
wenzelm 
parents: 
5802 
diff
changeset
 | 
70  | 
|
| 5737 | 71  | 
end  |