src/HOL/Induct/ABexp.thy
author paulson <lp15@cam.ac.uk>
Wed, 24 Apr 2024 20:56:26 +0100
changeset 80149 40a3fc07a587
parent 67443 3abf6a722518
permissions -rw-r--r--
More tidying of 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
    Author:     Stefan Berghofer, TU Muenchen
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     3
*)
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
     4
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     5
section \<open>Arithmetic and boolean expressions\<close>
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
     6
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 39246
diff changeset
     7
theory ABexp
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 39246
diff changeset
     8
imports Main
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 39246
diff changeset
     9
begin
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    10
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    11
datatype 'a aexp =
11046
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
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    21
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    22
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    23
text \<open>\medskip Evaluation of arithmetic and boolean expressions\<close>
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    24
60532
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    25
primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat"
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    26
  and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> bool"
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 39246
diff changeset
    27
where
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    28
  "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    29
| "evala env (Sum a1 a2) = evala env a1 + evala env a2"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    30
| "evala env (Diff a1 a2) = evala env a1 - evala env a2"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    31
| "evala env (Var v) = env v"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    32
| "evala env (Num n) = n"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    33
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    34
| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    35
| "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    36
| "evalb env (Neg b) = (\<not> evalb env b)"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    37
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    38
60530
44f9873d6f6f isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    39
text \<open>\medskip Substitution on arithmetic and boolean expressions\<close>
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    40
60532
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    41
primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
7fb5b7dc8332 misc tuning;
wenzelm
parents: 60530
diff changeset
    42
  and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp"
46914
c2ca2c3d23a6 misc tuning;
wenzelm
parents: 39246
diff changeset
    43
where
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    44
  "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    45
| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    46
| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    47
| "substa f (Var v) = f v"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    48
| "substa f (Num n) = Num n"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    49
39246
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    50
| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    51
| "substb f (And b1 b2) = And (substb f b1) (substb f b2)"
9e58f0499f57 modernized primrec
haftmann
parents: 36862
diff changeset
    52
| "substb f (Neg b) = Neg (substb f b)"
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    53
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    54
lemma subst1_aexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    55
  "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    56
and subst1_bexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    57
  "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63167
diff changeset
    58
    \<comment> \<open>one variable\<close>
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    59
  by (induct a and b) simp_all
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    60
12171
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    61
lemma subst_all_aexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    62
  "evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    63
and subst_all_bexp:
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    64
  "evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"
dc87f33db447 tuned inductions;
wenzelm
parents: 11649
diff changeset
    65
  by (induct a and b) auto
11046
b5f5942781a0 Induct: converted some theories to new-style format;
wenzelm
parents: 5802
diff changeset
    66
5737
31fc1d0e66d5 New example for using the datatype package:
berghofe
parents:
diff changeset
    67
end