author | Fabian Huch <huch@in.tum.de> |
Wed, 18 Oct 2023 20:26:02 +0200 | |
changeset 78844 | c7f436a63108 |
parent 76215 | a642599ffdea |
permissions | -rw-r--r-- |
9284 | 1 |
(* Title: ZF/Resid/Redex.thy |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
22808
diff
changeset
|
2 |
Author: Ole Rasmussen, University of Cambridge |
1048 | 3 |
*) |
4 |
||
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
61398
diff
changeset
|
5 |
theory Redex imports ZF begin |
1048 | 6 |
consts |
1401 | 7 |
redexes :: i |
1048 | 8 |
|
9 |
datatype |
|
12593 | 10 |
"redexes" = Var ("n \<in> nat") |
11 |
| Fun ("t \<in> redexes") |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12610
diff
changeset
|
12 |
| App ("b \<in> bool","f \<in> redexes", "a \<in> redexes") |
1048 | 13 |
|
6046 | 14 |
|
1048 | 15 |
consts |
12593 | 16 |
Ssub :: "i" |
17 |
Scomp :: "i" |
|
18 |
Sreg :: "i" |
|
22808 | 19 |
|
20 |
abbreviation |
|
69587 | 21 |
Ssub_rel (infixl \<open>\<Longleftarrow>\<close> 70) where |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
22 |
"a \<Longleftarrow> b \<equiv> \<langle>a,b\<rangle> \<in> Ssub" |
12593 | 23 |
|
22808 | 24 |
abbreviation |
69587 | 25 |
Scomp_rel (infixl \<open>\<sim>\<close> 70) where |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
26 |
"a \<sim> b \<equiv> \<langle>a,b\<rangle> \<in> Scomp" |
12593 | 27 |
|
22808 | 28 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
29 |
"regular(a) \<equiv> a \<in> Sreg" |
9284 | 30 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
31 |
consts union_aux :: "i\<Rightarrow>i" |
61398 | 32 |
primrec (*explicit lambda is required because both arguments of "\<squnion>" vary*) |
6046 | 33 |
"union_aux(Var(n)) = |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
34 |
(\<lambda>t \<in> redexes. redexes_case(\<lambda>j. Var(n), \<lambda>x. 0, \<lambda>b x y.0, t))" |
6046 | 35 |
|
36 |
"union_aux(Fun(u)) = |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
37 |
(\<lambda>t \<in> redexes. redexes_case(\<lambda>j. 0, \<lambda>y. Fun(union_aux(u)`y), |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
38 |
\<lambda>b y z. 0, t))" |
6046 | 39 |
|
40 |
"union_aux(App(b,f,a)) = |
|
12593 | 41 |
(\<lambda>t \<in> redexes. |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
42 |
redexes_case(\<lambda>j. 0, \<lambda>y. 0, |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76213
diff
changeset
|
43 |
\<lambda>c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))" |
6046 | 44 |
|
22808 | 45 |
definition |
69587 | 46 |
union (infixl \<open>\<squnion>\<close> 70) where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
47 |
"u \<squnion> v \<equiv> union_aux(u)`v" |
22808 | 48 |
|
6046 | 49 |
|
50 |
inductive |
|
46823 | 51 |
domains "Ssub" \<subseteq> "redexes*redexes" |
12593 | 52 |
intros |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
53 |
Sub_Var: "n \<in> nat \<Longrightarrow> Var(n) \<Longleftarrow> Var(n)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
54 |
Sub_Fun: "\<lbrakk>u \<Longleftarrow> v\<rbrakk>\<Longrightarrow> Fun(u) \<Longleftarrow> Fun(v)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
55 |
Sub_App1: "\<lbrakk>u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2; b \<in> bool\<rbrakk>\<Longrightarrow> |
61398 | 56 |
App(0,u1,u2) \<Longleftarrow> App(b,v1,v2)" |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
57 |
Sub_App2: "\<lbrakk>u1 \<Longleftarrow> v1; u2 \<Longleftarrow> v2\<rbrakk>\<Longrightarrow> App(1,u1,u2) \<Longleftarrow> App(1,v1,v2)" |
12593 | 58 |
type_intros redexes.intros bool_typechecks |
6046 | 59 |
|
60 |
inductive |
|
46823 | 61 |
domains "Scomp" \<subseteq> "redexes*redexes" |
12593 | 62 |
intros |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
63 |
Comp_Var: "n \<in> nat \<Longrightarrow> Var(n) \<sim> Var(n)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
64 |
Comp_Fun: "\<lbrakk>u \<sim> v\<rbrakk>\<Longrightarrow> Fun(u) \<sim> Fun(v)" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
65 |
Comp_App: "\<lbrakk>u1 \<sim> v1; u2 \<sim> v2; b1 \<in> bool; b2 \<in> bool\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
66 |
\<Longrightarrow> App(b1,u1,u2) \<sim> App(b2,v1,v2)" |
12593 | 67 |
type_intros redexes.intros bool_typechecks |
6046 | 68 |
|
69 |
inductive |
|
46823 | 70 |
domains "Sreg" \<subseteq> redexes |
12593 | 71 |
intros |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
72 |
Reg_Var: "n \<in> nat \<Longrightarrow> regular(Var(n))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
73 |
Reg_Fun: "\<lbrakk>regular(u)\<rbrakk>\<Longrightarrow> regular(Fun(u))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
74 |
Reg_App1: "\<lbrakk>regular(Fun(u)); regular(v)\<rbrakk>\<Longrightarrow>regular(App(1,Fun(u),v))" |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
75 |
Reg_App2: "\<lbrakk>regular(u); regular(v)\<rbrakk>\<Longrightarrow>regular(App(0,u,v))" |
12593 | 76 |
type_intros redexes.intros bool_typechecks |
77 |
||
78 |
||
79 |
declare redexes.intros [simp] |
|
80 |
||
81 |
(* ------------------------------------------------------------------------- *) |
|
82 |
(* Specialisation of comp-rules *) |
|
83 |
(* ------------------------------------------------------------------------- *) |
|
84 |
||
85 |
lemmas compD1 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD1] |
|
86 |
lemmas compD2 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD2] |
|
87 |
lemmas regD [simp] = Sreg.dom_subset [THEN subsetD] |
|
88 |
||
89 |
(* ------------------------------------------------------------------------- *) |
|
90 |
(* Equality rules for union *) |
|
91 |
(* ------------------------------------------------------------------------- *) |
|
92 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
93 |
lemma union_Var [simp]: "n \<in> nat \<Longrightarrow> Var(n) \<squnion> Var(n)=Var(n)" |
12593 | 94 |
by (simp add: union_def) |
95 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
96 |
lemma union_Fun [simp]: "v \<in> redexes \<Longrightarrow> Fun(u) \<squnion> Fun(v) = Fun(u \<squnion> v)" |
12593 | 97 |
by (simp add: union_def) |
98 |
||
99 |
lemma union_App [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
100 |
"\<lbrakk>b2 \<in> bool; u2 \<in> redexes; v2 \<in> redexes\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
101 |
\<Longrightarrow> App(b1,u1,v1) \<squnion> App(b2,u2,v2)=App(b1 or b2,u1 \<squnion> u2,v1 \<squnion> v2)" |
12593 | 102 |
by (simp add: union_def) |
103 |
||
104 |
||
105 |
lemma or_1_right [simp]: "a or 1 = 1" |
|
106 |
by (simp add: or_def cond_def) |
|
107 |
||
108 |
lemma or_0_right [simp]: "a \<in> bool \<Longrightarrow> a or 0 = a" |
|
109 |
by (simp add: or_def cond_def bool_def, auto) |
|
110 |
||
111 |
||
112 |
||
113 |
declare Ssub.intros [simp] |
|
114 |
declare bool_typechecks [simp] |
|
115 |
declare Sreg.intros [simp] |
|
116 |
declare Scomp.intros [simp] |
|
117 |
||
118 |
declare Scomp.intros [intro] |
|
6046 | 119 |
|
12610 | 120 |
inductive_cases [elim!]: |
121 |
"regular(App(b,f,a))" |
|
122 |
"regular(Fun(b))" |
|
123 |
"regular(Var(b))" |
|
61398 | 124 |
"Fun(u) \<sim> Fun(t)" |
125 |
"u \<sim> Fun(t)" |
|
126 |
"u \<sim> Var(n)" |
|
127 |
"u \<sim> App(b,t,a)" |
|
128 |
"Fun(t) \<sim> v" |
|
129 |
"App(b,f,a) \<sim> v" |
|
130 |
"Var(n) \<sim> u" |
|
12593 | 131 |
|
132 |
||
133 |
||
134 |
(* ------------------------------------------------------------------------- *) |
|
135 |
(* comp proofs *) |
|
136 |
(* ------------------------------------------------------------------------- *) |
|
137 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
138 |
lemma comp_refl [simp]: "u \<in> redexes \<Longrightarrow> u \<sim> u" |
12593 | 139 |
by (erule redexes.induct, blast+) |
140 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
141 |
lemma comp_sym: "u \<sim> v \<Longrightarrow> v \<sim> u" |
12593 | 142 |
by (erule Scomp.induct, blast+) |
143 |
||
61398 | 144 |
lemma comp_sym_iff: "u \<sim> v \<longleftrightarrow> v \<sim> u" |
12593 | 145 |
by (blast intro: comp_sym) |
146 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69587
diff
changeset
|
147 |
lemma comp_trans [rule_format]: "u \<sim> v \<Longrightarrow> \<forall>w. v \<sim> w\<longrightarrow>u \<sim> w" |
12593 | 148 |
by (erule Scomp.induct, blast+) |
149 |
||
150 |
(* ------------------------------------------------------------------------- *) |
|
151 |
(* union proofs *) |
|
152 |
(* ------------------------------------------------------------------------- *) |
|
153 |
||
61398 | 154 |
lemma union_l: "u \<sim> v \<Longrightarrow> u \<Longleftarrow> (u \<squnion> v)" |
12593 | 155 |
apply (erule Scomp.induct) |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12610
diff
changeset
|
156 |
apply (erule_tac [3] boolE, simp_all) |
12593 | 157 |
done |
158 |
||
61398 | 159 |
lemma union_r: "u \<sim> v \<Longrightarrow> v \<Longleftarrow> (u \<squnion> v)" |
12593 | 160 |
apply (erule Scomp.induct) |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12610
diff
changeset
|
161 |
apply (erule_tac [3] c = b2 in boolE, simp_all) |
12593 | 162 |
done |
163 |
||
61398 | 164 |
lemma union_sym: "u \<sim> v \<Longrightarrow> u \<squnion> v = v \<squnion> u" |
12593 | 165 |
by (erule Scomp.induct, simp_all add: or_commute) |
166 |
||
167 |
(* ------------------------------------------------------------------------- *) |
|
168 |
(* regular proofs *) |
|
169 |
(* ------------------------------------------------------------------------- *) |
|
170 |
||
171 |
lemma union_preserve_regular [rule_format]: |
|
61398 | 172 |
"u \<sim> v \<Longrightarrow> regular(u) \<longrightarrow> regular(v) \<longrightarrow> regular(u \<squnion> v)" |
13612 | 173 |
by (erule Scomp.induct, auto) |
6046 | 174 |
|
1048 | 175 |
end |
176 |