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