author | Lukas Bartl |
Thu, 09 Jan 2025 13:18:37 +0100 | |
changeset 81757 | 4d15005da582 |
parent 76215 | a642599ffdea |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/IMP/Com.thy |
40945 | 2 |
Author: Heiko Loetzbeyer and Robert Sandner, TU München |
12610 | 3 |
*) |
482 | 4 |
|
60770 | 5 |
section \<open>Arithmetic expressions, boolean expressions, commands\<close> |
482 | 6 |
|
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
60770
diff
changeset
|
7 |
theory Com imports ZF begin |
511 | 8 |
|
12610 | 9 |
|
60770 | 10 |
subsection \<open>Arithmetic expressions\<close> |
511 | 11 |
|
12610 | 12 |
consts |
13 |
loc :: i |
|
14 |
aexp :: i |
|
15 |
||
16 |
datatype \<subseteq> "univ(loc \<union> (nat -> nat) \<union> ((nat \<times> nat) -> nat))" |
|
17 |
aexp = N ("n \<in> nat") |
|
18 |
| X ("x \<in> loc") |
|
19 |
| Op1 ("f \<in> nat -> nat", "a \<in> aexp") |
|
20 |
| Op2 ("f \<in> (nat \<times> nat) -> nat", "a0 \<in> aexp", "a1 \<in> aexp") |
|
511 | 21 |
|
22 |
||
12610 | 23 |
consts evala :: i |
35068 | 24 |
|
25 |
abbreviation |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
26 |
evala_syntax :: "[i, i] \<Rightarrow> o" (infixl \<open>-a->\<close> 50) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
27 |
where "p -a-> n \<equiv> \<langle>p,n\<rangle> \<in> evala" |
12610 | 28 |
|
511 | 29 |
inductive |
12610 | 30 |
domains "evala" \<subseteq> "(aexp \<times> (loc -> nat)) \<times> nat" |
31 |
intros |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
32 |
N: "\<lbrakk>n \<in> nat; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <N(n),sigma> -a-> n" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
33 |
X: "\<lbrakk>x \<in> loc; sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <X(x),sigma> -a-> sigma`x" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
34 |
Op1: "\<lbrakk>\<langle>e,sigma\<rangle> -a-> n; f \<in> nat -> nat\<rbrakk> \<Longrightarrow> <Op1(f,e),sigma> -a-> f`n" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
35 |
Op2: "\<lbrakk>\<langle>e0,sigma\<rangle> -a-> n0; \<langle>e1,sigma\<rangle> -a-> n1; f \<in> (nat\<times>nat) -> nat\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
36 |
\<Longrightarrow> <Op2(f,e0,e1),sigma> -a-> f`\<langle>n0,n1\<rangle>" |
12606 | 37 |
type_intros aexp.intros apply_funtype |
511 | 38 |
|
39 |
||
60770 | 40 |
subsection \<open>Boolean expressions\<close> |
12610 | 41 |
|
42 |
consts bexp :: i |
|
511 | 43 |
|
12610 | 44 |
datatype \<subseteq> "univ(aexp \<union> ((nat \<times> nat)->bool))" |
45 |
bexp = true |
|
46 |
| false |
|
47 |
| ROp ("f \<in> (nat \<times> nat)->bool", "a0 \<in> aexp", "a1 \<in> aexp") |
|
48 |
| noti ("b \<in> bexp") |
|
69587 | 49 |
| andi ("b0 \<in> bexp", "b1 \<in> bexp") (infixl \<open>andi\<close> 60) |
50 |
| ori ("b0 \<in> bexp", "b1 \<in> bexp") (infixl \<open>ori\<close> 60) |
|
511 | 51 |
|
52 |
||
12610 | 53 |
consts evalb :: i |
35068 | 54 |
|
55 |
abbreviation |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
56 |
evalb_syntax :: "[i,i] \<Rightarrow> o" (infixl \<open>-b->\<close> 50) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
57 |
where "p -b-> b \<equiv> \<langle>p,b\<rangle> \<in> evalb" |
511 | 58 |
|
59 |
inductive |
|
12610 | 60 |
domains "evalb" \<subseteq> "(bexp \<times> (loc -> nat)) \<times> bool" |
61 |
intros |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
62 |
true: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> \<langle>true,sigma\<rangle> -b-> 1" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
63 |
false: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> \<langle>false,sigma\<rangle> -b-> 0" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
64 |
ROp: "\<lbrakk>\<langle>a0,sigma\<rangle> -a-> n0; \<langle>a1,sigma\<rangle> -a-> n1; f \<in> (nat*nat)->bool\<rbrakk> |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
65 |
\<Longrightarrow> <ROp(f,a0,a1),sigma> -b-> f`\<langle>n0,n1\<rangle> " |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
66 |
noti: "\<lbrakk>\<langle>b,sigma\<rangle> -b-> w\<rbrakk> \<Longrightarrow> <noti(b),sigma> -b-> not(w)" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
67 |
andi: "\<lbrakk>\<langle>b0,sigma\<rangle> -b-> w0; \<langle>b1,sigma\<rangle> -b-> w1\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
68 |
\<Longrightarrow> <b0 andi b1,sigma> -b-> (w0 and w1)" |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
69 |
ori: "\<lbrakk>\<langle>b0,sigma\<rangle> -b-> w0; \<langle>b1,sigma\<rangle> -b-> w1\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
70 |
\<Longrightarrow> <b0 ori b1,sigma> -b-> (w0 or w1)" |
12610 | 71 |
type_intros bexp.intros |
12606 | 72 |
apply_funtype and_type or_type bool_1I bool_0I not_type |
73 |
type_elims evala.dom_subset [THEN subsetD, elim_format] |
|
511 | 74 |
|
75 |
||
60770 | 76 |
subsection \<open>Commands\<close> |
511 | 77 |
|
12610 | 78 |
consts com :: i |
79 |
datatype com = |
|
69587 | 80 |
skip (\<open>\<SKIP>\<close> []) |
81 |
| assignment ("x \<in> loc", "a \<in> aexp") (infixl \<open>\<ASSN>\<close> 60) |
|
82 |
| semicolon ("c0 \<in> com", "c1 \<in> com") (\<open>_\<SEQ> _\<close> [60, 60] 10) |
|
83 |
| while ("b \<in> bexp", "c \<in> com") (\<open>\<WHILE> _ \<DO> _\<close> 60) |
|
84 |
| "if" ("b \<in> bexp", "c0 \<in> com", "c1 \<in> com") (\<open>\<IF> _ \<THEN> _ \<ELSE> _\<close> 60) |
|
511 | 85 |
|
86 |
||
12610 | 87 |
consts evalc :: i |
35068 | 88 |
|
89 |
abbreviation |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
90 |
evalc_syntax :: "[i, i] \<Rightarrow> o" (infixl \<open>-c->\<close> 50) |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
91 |
where "p -c-> s \<equiv> \<langle>p,s\<rangle> \<in> evalc" |
511 | 92 |
|
93 |
inductive |
|
12610 | 94 |
domains "evalc" \<subseteq> "(com \<times> (loc -> nat)) \<times> (loc -> nat)" |
12606 | 95 |
intros |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
96 |
skip: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <\<SKIP>,sigma> -c-> sigma" |
12606 | 97 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
98 |
assign: "\<lbrakk>m \<in> nat; x \<in> loc; \<langle>a,sigma\<rangle> -a-> m\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
99 |
\<Longrightarrow> <x \<ASSN> a,sigma> -c-> sigma(x:=m)" |
511 | 100 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
101 |
semi: "\<lbrakk>\<langle>c0,sigma\<rangle> -c-> sigma2; \<langle>c1,sigma2\<rangle> -c-> sigma1\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
102 |
\<Longrightarrow> <c0\<SEQ> c1, sigma> -c-> sigma1" |
12606 | 103 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
104 |
if1: "\<lbrakk>b \<in> bexp; c1 \<in> com; sigma \<in> loc->nat; |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
105 |
\<langle>b,sigma\<rangle> -b-> 1; \<langle>c0,sigma\<rangle> -c-> sigma1\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
106 |
\<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1" |
511 | 107 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
108 |
if0: "\<lbrakk>b \<in> bexp; c0 \<in> com; sigma \<in> loc->nat; |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
109 |
\<langle>b,sigma\<rangle> -b-> 0; \<langle>c1,sigma\<rangle> -c-> sigma1\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
110 |
\<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1" |
12606 | 111 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
112 |
while0: "\<lbrakk>c \<in> com; \<langle>b, sigma\<rangle> -b-> 0\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
113 |
\<Longrightarrow> <\<WHILE> b \<DO> c,sigma> -c-> sigma" |
511 | 114 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
115 |
while1: "\<lbrakk>c \<in> com; \<langle>b,sigma\<rangle> -b-> 1; \<langle>c,sigma\<rangle> -c-> sigma2; |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
116 |
<\<WHILE> b \<DO> c, sigma2> -c-> sigma1\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
117 |
\<Longrightarrow> <\<WHILE> b \<DO> c, sigma> -c-> sigma1" |
12606 | 118 |
|
119 |
type_intros com.intros update_type |
|
120 |
type_elims evala.dom_subset [THEN subsetD, elim_format] |
|
121 |
evalb.dom_subset [THEN subsetD, elim_format] |
|
122 |
||
123 |
||
60770 | 124 |
subsection \<open>Misc lemmas\<close> |
511 | 125 |
|
12610 | 126 |
lemmas evala_1 [simp] = evala.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD1] |
127 |
and evala_2 [simp] = evala.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD2] |
|
128 |
and evala_3 [simp] = evala.dom_subset [THEN subsetD, THEN SigmaD2] |
|
511 | 129 |
|
12610 | 130 |
lemmas evalb_1 [simp] = evalb.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD1] |
131 |
and evalb_2 [simp] = evalb.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD2] |
|
132 |
and evalb_3 [simp] = evalb.dom_subset [THEN subsetD, THEN SigmaD2] |
|
511 | 133 |
|
12610 | 134 |
lemmas evalc_1 [simp] = evalc.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD1] |
135 |
and evalc_2 [simp] = evalc.dom_subset [THEN subsetD, THEN SigmaD1, THEN SigmaD2] |
|
136 |
and evalc_3 [simp] = evalc.dom_subset [THEN subsetD, THEN SigmaD2] |
|
511 | 137 |
|
12610 | 138 |
inductive_cases |
139 |
evala_N_E [elim!]: "<N(n),sigma> -a-> i" |
|
140 |
and evala_X_E [elim!]: "<X(x),sigma> -a-> i" |
|
141 |
and evala_Op1_E [elim!]: "<Op1(f,e),sigma> -a-> i" |
|
142 |
and evala_Op2_E [elim!]: "<Op2(f,a1,a2),sigma> -a-> i" |
|
511 | 143 |
|
144 |
end |