author | wenzelm |
Sun, 05 Jun 2022 20:14:32 +0200 | |
changeset 75513 | 36316c6a3fc2 |
parent 69593 | 3dda49e08b9d |
child 76213 | e44d86131648 |
permissions | -rw-r--r-- |
12560 | 1 |
(* Title: ZF/Induct/Comb.thy |
12088 | 2 |
Author: Lawrence C Paulson |
3 |
Copyright 1994 University of Cambridge |
|
4 |
*) |
|
5 |
||
60770 | 6 |
section \<open>Combinatory Logic example: the Church-Rosser Theorem\<close> |
12088 | 7 |
|
62045 | 8 |
theory Comb |
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
62045
diff
changeset
|
9 |
imports ZF |
62045 | 10 |
begin |
12088 | 11 |
|
60770 | 12 |
text \<open> |
12610 | 13 |
Curiously, combinators do not include free variables. |
14 |
||
58623 | 15 |
Example taken from @{cite camilleri92}. |
60770 | 16 |
\<close> |
12610 | 17 |
|
62045 | 18 |
|
60770 | 19 |
subsection \<open>Definitions\<close> |
12610 | 20 |
|
61798 | 21 |
text \<open>Datatype definition of combinators \<open>S\<close> and \<open>K\<close>.\<close> |
12088 | 22 |
|
12610 | 23 |
consts comb :: i |
24 |
datatype comb = |
|
62045 | 25 |
K |
26 |
| S |
|
69587 | 27 |
| app ("p \<in> comb", "q \<in> comb") (infixl \<open>\<bullet>\<close> 90) |
35427 | 28 |
|
60770 | 29 |
text \<open> |
62045 | 30 |
Inductive definition of contractions, \<open>\<rightarrow>\<^sup>1\<close> and |
31 |
(multi-step) reductions, \<open>\<rightarrow>\<close>. |
|
60770 | 32 |
\<close> |
12610 | 33 |
|
62045 | 34 |
consts contract :: i |
69587 | 35 |
abbreviation contract_syntax :: "[i,i] \<Rightarrow> o" (infixl \<open>\<rightarrow>\<^sup>1\<close> 50) |
62045 | 36 |
where "p \<rightarrow>\<^sup>1 q \<equiv> <p,q> \<in> contract" |
35068 | 37 |
|
69587 | 38 |
abbreviation contract_multi :: "[i,i] \<Rightarrow> o" (infixl \<open>\<rightarrow>\<close> 50) |
62045 | 39 |
where "p \<rightarrow> q \<equiv> <p,q> \<in> contract^*" |
12088 | 40 |
|
41 |
inductive |
|
12610 | 42 |
domains "contract" \<subseteq> "comb \<times> comb" |
12560 | 43 |
intros |
62045 | 44 |
K: "[| p \<in> comb; q \<in> comb |] ==> K\<bullet>p\<bullet>q \<rightarrow>\<^sup>1 p" |
45 |
S: "[| p \<in> comb; q \<in> comb; r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r \<rightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)" |
|
46 |
Ap1: "[| p\<rightarrow>\<^sup>1q; r \<in> comb |] ==> p\<bullet>r \<rightarrow>\<^sup>1 q\<bullet>r" |
|
47 |
Ap2: "[| p\<rightarrow>\<^sup>1q; r \<in> comb |] ==> r\<bullet>p \<rightarrow>\<^sup>1 r\<bullet>q" |
|
12610 | 48 |
type_intros comb.intros |
12088 | 49 |
|
60770 | 50 |
text \<open> |
62045 | 51 |
Inductive definition of parallel contractions, \<open>\<Rrightarrow>\<^sup>1\<close> and |
52 |
(multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>. |
|
60770 | 53 |
\<close> |
12610 | 54 |
|
62045 | 55 |
consts parcontract :: i |
35068 | 56 |
|
69587 | 57 |
abbreviation parcontract_syntax :: "[i,i] => o" (infixl \<open>\<Rrightarrow>\<^sup>1\<close> 50) |
62045 | 58 |
where "p \<Rrightarrow>\<^sup>1 q == <p,q> \<in> parcontract" |
35068 | 59 |
|
69587 | 60 |
abbreviation parcontract_multi :: "[i,i] => o" (infixl \<open>\<Rrightarrow>\<close> 50) |
62045 | 61 |
where "p \<Rrightarrow> q == <p,q> \<in> parcontract^+" |
12088 | 62 |
|
63 |
inductive |
|
12610 | 64 |
domains "parcontract" \<subseteq> "comb \<times> comb" |
12560 | 65 |
intros |
62045 | 66 |
refl: "[| p \<in> comb |] ==> p \<Rrightarrow>\<^sup>1 p" |
67 |
K: "[| p \<in> comb; q \<in> comb |] ==> K\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 p" |
|
68 |
S: "[| p \<in> comb; q \<in> comb; r \<in> comb |] ==> S\<bullet>p\<bullet>q\<bullet>r \<Rrightarrow>\<^sup>1 (p\<bullet>r)\<bullet>(q\<bullet>r)" |
|
69 |
Ap: "[| p\<Rrightarrow>\<^sup>1q; r\<Rrightarrow>\<^sup>1s |] ==> p\<bullet>r \<Rrightarrow>\<^sup>1 q\<bullet>s" |
|
12610 | 70 |
type_intros comb.intros |
12088 | 71 |
|
60770 | 72 |
text \<open> |
12610 | 73 |
Misc definitions. |
60770 | 74 |
\<close> |
12088 | 75 |
|
62045 | 76 |
definition I :: i |
77 |
where "I \<equiv> S\<bullet>K\<bullet>K" |
|
12088 | 78 |
|
62045 | 79 |
definition diamond :: "i \<Rightarrow> o" |
80 |
where "diamond(r) \<equiv> |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35427
diff
changeset
|
81 |
\<forall>x y. <x,y>\<in>r \<longrightarrow> (\<forall>y'. <x,y'>\<in>r \<longrightarrow> (\<exists>z. <y,z>\<in>r & <y',z> \<in> r))" |
12560 | 82 |
|
83 |
||
60770 | 84 |
subsection \<open>Transitive closure preserves the Church-Rosser property\<close> |
12560 | 85 |
|
86 |
lemma diamond_strip_lemmaD [rule_format]: |
|
12610 | 87 |
"[| diamond(r); <x,y>:r^+ |] ==> |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
35427
diff
changeset
|
88 |
\<forall>y'. <x,y'>:r \<longrightarrow> (\<exists>z. <y',z>: r^+ & <y,z>: r)" |
12610 | 89 |
apply (unfold diamond_def) |
90 |
apply (erule trancl_induct) |
|
91 |
apply (blast intro: r_into_trancl) |
|
92 |
apply clarify |
|
93 |
apply (drule spec [THEN mp], assumption) |
|
94 |
apply (blast intro: r_into_trancl trans_trancl [THEN transD]) |
|
95 |
done |
|
12560 | 96 |
|
97 |
lemma diamond_trancl: "diamond(r) ==> diamond(r^+)" |
|
12610 | 98 |
apply (simp (no_asm_simp) add: diamond_def) |
99 |
apply (rule impI [THEN allI, THEN allI]) |
|
100 |
apply (erule trancl_induct) |
|
101 |
apply auto |
|
102 |
apply (best intro: r_into_trancl trans_trancl [THEN transD] |
|
103 |
dest: diamond_strip_lemmaD)+ |
|
104 |
done |
|
12560 | 105 |
|
15775 | 106 |
inductive_cases Ap_E [elim!]: "p\<bullet>q \<in> comb" |
12560 | 107 |
|
108 |
||
60770 | 109 |
subsection \<open>Results about Contraction\<close> |
12560 | 110 |
|
60770 | 111 |
text \<open> |
69593 | 112 |
For type checking: replaces \<^term>\<open>a \<rightarrow>\<^sup>1 b\<close> by \<open>a, b \<in> |
61798 | 113 |
comb\<close>. |
60770 | 114 |
\<close> |
12610 | 115 |
|
12560 | 116 |
lemmas contract_combE2 = contract.dom_subset [THEN subsetD, THEN SigmaE2] |
12610 | 117 |
and contract_combD1 = contract.dom_subset [THEN subsetD, THEN SigmaD1] |
118 |
and contract_combD2 = contract.dom_subset [THEN subsetD, THEN SigmaD2] |
|
12560 | 119 |
|
120 |
lemma field_contract_eq: "field(contract) = comb" |
|
12610 | 121 |
by (blast intro: contract.K elim!: contract_combE2) |
12560 | 122 |
|
12610 | 123 |
lemmas reduction_refl = |
124 |
field_contract_eq [THEN equalityD2, THEN subsetD, THEN rtrancl_refl] |
|
12560 | 125 |
|
12610 | 126 |
lemmas rtrancl_into_rtrancl2 = |
127 |
r_into_rtrancl [THEN trans_rtrancl [THEN transD]] |
|
12560 | 128 |
|
129 |
declare reduction_refl [intro!] contract.K [intro!] contract.S [intro!] |
|
130 |
||
12610 | 131 |
lemmas reduction_rls = |
132 |
contract.K [THEN rtrancl_into_rtrancl2] |
|
133 |
contract.S [THEN rtrancl_into_rtrancl2] |
|
134 |
contract.Ap1 [THEN rtrancl_into_rtrancl2] |
|
135 |
contract.Ap2 [THEN rtrancl_into_rtrancl2] |
|
12560 | 136 |
|
62045 | 137 |
lemma "p \<in> comb ==> I\<bullet>p \<rightarrow> p" |
61798 | 138 |
\<comment> \<open>Example only: not used\<close> |
62045 | 139 |
unfolding I_def by (blast intro: reduction_rls) |
12560 | 140 |
|
141 |
lemma comb_I: "I \<in> comb" |
|
62045 | 142 |
unfolding I_def by blast |
12560 | 143 |
|
12610 | 144 |
|
60770 | 145 |
subsection \<open>Non-contraction results\<close> |
12560 | 146 |
|
60770 | 147 |
text \<open>Derive a case for each combinator constructor.\<close> |
12610 | 148 |
|
62045 | 149 |
inductive_cases K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r" |
150 |
and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r" |
|
151 |
and Ap_contractE [elim!]: "p\<bullet>q \<rightarrow>\<^sup>1 r" |
|
12560 | 152 |
|
62045 | 153 |
lemma I_contract_E: "I \<rightarrow>\<^sup>1 r ==> P" |
12610 | 154 |
by (auto simp add: I_def) |
12560 | 155 |
|
62045 | 156 |
lemma K1_contractD: "K\<bullet>p \<rightarrow>\<^sup>1 r ==> (\<exists>q. r = K\<bullet>q & p \<rightarrow>\<^sup>1 q)" |
12610 | 157 |
by auto |
12560 | 158 |
|
62045 | 159 |
lemma Ap_reduce1: "[| p \<rightarrow> q; r \<in> comb |] ==> p\<bullet>r \<rightarrow> q\<bullet>r" |
12610 | 160 |
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) |
161 |
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) |
|
162 |
apply (erule rtrancl_induct) |
|
163 |
apply (blast intro: reduction_rls) |
|
164 |
apply (erule trans_rtrancl [THEN transD]) |
|
165 |
apply (blast intro: contract_combD2 reduction_rls) |
|
166 |
done |
|
12560 | 167 |
|
62045 | 168 |
lemma Ap_reduce2: "[| p \<rightarrow> q; r \<in> comb |] ==> r\<bullet>p \<rightarrow> r\<bullet>q" |
12610 | 169 |
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) |
170 |
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) |
|
171 |
apply (erule rtrancl_induct) |
|
172 |
apply (blast intro: reduction_rls) |
|
13573 | 173 |
apply (blast intro: trans_rtrancl [THEN transD] |
174 |
contract_combD2 reduction_rls) |
|
12610 | 175 |
done |
12560 | 176 |
|
62045 | 177 |
text \<open>Counterexample to the diamond property for \<open>\<rightarrow>\<^sup>1\<close>.\<close> |
12560 | 178 |
|
62045 | 179 |
lemma KIII_contract1: "K\<bullet>I\<bullet>(I\<bullet>I) \<rightarrow>\<^sup>1 I" |
46900 | 180 |
by (blast intro: comb_I) |
12560 | 181 |
|
62045 | 182 |
lemma KIII_contract2: "K\<bullet>I\<bullet>(I\<bullet>I) \<rightarrow>\<^sup>1 K\<bullet>I\<bullet>((K\<bullet>I)\<bullet>(K\<bullet>I))" |
46900 | 183 |
by (unfold I_def) (blast intro: contract.intros) |
12560 | 184 |
|
62045 | 185 |
lemma KIII_contract3: "K\<bullet>I\<bullet>((K\<bullet>I)\<bullet>(K\<bullet>I)) \<rightarrow>\<^sup>1 I" |
46900 | 186 |
by (blast intro: comb_I) |
12560 | 187 |
|
12610 | 188 |
lemma not_diamond_contract: "\<not> diamond(contract)" |
189 |
apply (unfold diamond_def) |
|
190 |
apply (blast intro: KIII_contract1 KIII_contract2 KIII_contract3 |
|
191 |
elim!: I_contract_E) |
|
192 |
done |
|
12560 | 193 |
|
194 |
||
60770 | 195 |
subsection \<open>Results about Parallel Contraction\<close> |
12560 | 196 |
|
62045 | 197 |
text \<open>For type checking: replaces \<open>a \<Rrightarrow>\<^sup>1 b\<close> by \<open>a, b |
61798 | 198 |
\<in> comb\<close>\<close> |
12560 | 199 |
lemmas parcontract_combE2 = parcontract.dom_subset [THEN subsetD, THEN SigmaE2] |
12610 | 200 |
and parcontract_combD1 = parcontract.dom_subset [THEN subsetD, THEN SigmaD1] |
201 |
and parcontract_combD2 = parcontract.dom_subset [THEN subsetD, THEN SigmaD2] |
|
12560 | 202 |
|
203 |
lemma field_parcontract_eq: "field(parcontract) = comb" |
|
12610 | 204 |
by (blast intro: parcontract.K elim!: parcontract_combE2) |
12560 | 205 |
|
60770 | 206 |
text \<open>Derive a case for each combinator constructor.\<close> |
12610 | 207 |
inductive_cases |
62045 | 208 |
K_parcontractE [elim!]: "K \<Rrightarrow>\<^sup>1 r" |
209 |
and S_parcontractE [elim!]: "S \<Rrightarrow>\<^sup>1 r" |
|
210 |
and Ap_parcontractE [elim!]: "p\<bullet>q \<Rrightarrow>\<^sup>1 r" |
|
12560 | 211 |
|
212 |
declare parcontract.intros [intro] |
|
213 |
||
214 |
||
60770 | 215 |
subsection \<open>Basic properties of parallel contraction\<close> |
12560 | 216 |
|
12610 | 217 |
lemma K1_parcontractD [dest!]: |
62045 | 218 |
"K\<bullet>p \<Rrightarrow>\<^sup>1 r ==> (\<exists>p'. r = K\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')" |
12610 | 219 |
by auto |
12560 | 220 |
|
12610 | 221 |
lemma S1_parcontractD [dest!]: |
62045 | 222 |
"S\<bullet>p \<Rrightarrow>\<^sup>1 r ==> (\<exists>p'. r = S\<bullet>p' & p \<Rrightarrow>\<^sup>1 p')" |
12610 | 223 |
by auto |
12560 | 224 |
|
225 |
lemma S2_parcontractD [dest!]: |
|
62045 | 226 |
"S\<bullet>p\<bullet>q \<Rrightarrow>\<^sup>1 r ==> (\<exists>p' q'. r = S\<bullet>p'\<bullet>q' & p \<Rrightarrow>\<^sup>1 p' & q \<Rrightarrow>\<^sup>1 q')" |
12610 | 227 |
by auto |
12560 | 228 |
|
229 |
lemma diamond_parcontract: "diamond(parcontract)" |
|
61798 | 230 |
\<comment> \<open>Church-Rosser property for parallel contraction\<close> |
12610 | 231 |
apply (unfold diamond_def) |
232 |
apply (rule impI [THEN allI, THEN allI]) |
|
233 |
apply (erule parcontract.induct) |
|
234 |
apply (blast elim!: comb.free_elims intro: parcontract_combD2)+ |
|
235 |
done |
|
12560 | 236 |
|
60770 | 237 |
text \<open> |
69593 | 238 |
\medskip Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>. |
60770 | 239 |
\<close> |
12560 | 240 |
|
62045 | 241 |
lemma contract_imp_parcontract: "p\<rightarrow>\<^sup>1q ==> p\<Rrightarrow>\<^sup>1q" |
18415 | 242 |
by (induct set: contract) auto |
12560 | 243 |
|
62045 | 244 |
lemma reduce_imp_parreduce: "p\<rightarrow>q ==> p\<Rrightarrow>q" |
12610 | 245 |
apply (frule rtrancl_type [THEN subsetD, THEN SigmaD1]) |
246 |
apply (drule field_contract_eq [THEN equalityD1, THEN subsetD]) |
|
247 |
apply (erule rtrancl_induct) |
|
248 |
apply (blast intro: r_into_trancl) |
|
249 |
apply (blast intro: contract_imp_parcontract r_into_trancl |
|
250 |
trans_trancl [THEN transD]) |
|
251 |
done |
|
12560 | 252 |
|
62045 | 253 |
lemma parcontract_imp_reduce: "p\<Rrightarrow>\<^sup>1q ==> p\<rightarrow>q" |
18415 | 254 |
apply (induct set: parcontract) |
12610 | 255 |
apply (blast intro: reduction_rls) |
256 |
apply (blast intro: reduction_rls) |
|
12560 | 257 |
apply (blast intro: reduction_rls) |
12610 | 258 |
apply (blast intro: trans_rtrancl [THEN transD] |
259 |
Ap_reduce1 Ap_reduce2 parcontract_combD1 parcontract_combD2) |
|
260 |
done |
|
12560 | 261 |
|
62045 | 262 |
lemma parreduce_imp_reduce: "p\<Rrightarrow>q ==> p\<rightarrow>q" |
12610 | 263 |
apply (frule trancl_type [THEN subsetD, THEN SigmaD1]) |
264 |
apply (drule field_parcontract_eq [THEN equalityD1, THEN subsetD]) |
|
265 |
apply (erule trancl_induct, erule parcontract_imp_reduce) |
|
266 |
apply (erule trans_rtrancl [THEN transD]) |
|
267 |
apply (erule parcontract_imp_reduce) |
|
268 |
done |
|
12560 | 269 |
|
62045 | 270 |
lemma parreduce_iff_reduce: "p\<Rrightarrow>q \<longleftrightarrow> p\<rightarrow>q" |
12610 | 271 |
by (blast intro: parreduce_imp_reduce reduce_imp_parreduce) |
12088 | 272 |
|
273 |
end |