author | wenzelm |
Wed, 28 Dec 2022 12:30:18 +0100 | |
changeset 76798 | 69d8d16c5612 |
parent 71591 | 8e4d542f041b |
child 76987 | 4c275405faae |
permissions | -rw-r--r-- |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
1 |
(* Title: HOL/Induct/Comb.thy |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
3 |
Copyright 1996 University of Cambridge |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
4 |
*) |
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
5 |
|
60530 | 6 |
section \<open>Combinatory Logic example: the Church-Rosser Theorem\<close> |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
7 |
|
62045 | 8 |
theory Comb |
9 |
imports Main |
|
10 |
begin |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
11 |
|
60530 | 12 |
text \<open> |
71591 | 13 |
Combinator terms do not have free variables. |
58623 | 14 |
Example taken from @{cite camilleri92}. |
62045 | 15 |
\<close> |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
16 |
|
60530 | 17 |
subsection \<open>Definitions\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
18 |
|
61993 | 19 |
text \<open>Datatype definition of combinators \<open>S\<close> and \<open>K\<close>.\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
20 |
|
58310 | 21 |
datatype comb = K |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
22 |
| S |
61993 | 23 |
| Ap comb comb (infixl "\<bullet>" 90) |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
24 |
|
60530 | 25 |
text \<open> |
61993 | 26 |
Inductive definition of contractions, \<open>\<rightarrow>\<^sup>1\<close> and |
27 |
(multi-step) reductions, \<open>\<rightarrow>\<close>. |
|
60530 | 28 |
\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
29 |
|
71591 | 30 |
inductive contract1 :: "[comb,comb] \<Rightarrow> bool" (infixl "\<rightarrow>\<^sup>1" 50) |
31 |
where |
|
32 |
K: "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x" |
|
33 |
| S: "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)" |
|
34 |
| Ap1: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x\<bullet>z \<rightarrow>\<^sup>1 y\<bullet>z" |
|
35 |
| Ap2: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y" |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
36 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
37 |
abbreviation |
71591 | 38 |
contract :: "[comb,comb] \<Rightarrow> bool" (infixl "\<rightarrow>" 50) where |
39 |
"contract \<equiv> contract1\<^sup>*\<^sup>*" |
|
15816 | 40 |
|
60530 | 41 |
text \<open> |
61993 | 42 |
Inductive definition of parallel contractions, \<open>\<Rrightarrow>\<^sup>1\<close> and |
43 |
(multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>. |
|
60530 | 44 |
\<close> |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
45 |
|
71591 | 46 |
inductive parcontract1 :: "[comb,comb] \<Rightarrow> bool" (infixl "\<Rrightarrow>\<^sup>1" 50) |
47 |
where |
|
48 |
refl: "x \<Rrightarrow>\<^sup>1 x" |
|
49 |
| K: "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x" |
|
50 |
| S: "S\<bullet>x\<bullet>y\<bullet>z \<Rrightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)" |
|
51 |
| Ap: "\<lbrakk>x \<Rrightarrow>\<^sup>1 y; z \<Rrightarrow>\<^sup>1 w\<rbrakk> \<Longrightarrow> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w" |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
52 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
53 |
abbreviation |
71591 | 54 |
parcontract :: "[comb,comb] \<Rightarrow> bool" (infixl "\<Rrightarrow>" 50) where |
55 |
"parcontract \<equiv> parcontract1\<^sup>*\<^sup>*" |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
56 |
|
60530 | 57 |
text \<open> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
58 |
Misc definitions. |
60530 | 59 |
\<close> |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
60 |
|
19736 | 61 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
62 |
I :: comb where |
71591 | 63 |
"I \<equiv> S\<bullet>K\<bullet>K" |
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
64 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
65 |
definition |
71591 | 66 |
diamond :: "([comb,comb] \<Rightarrow> bool) \<Rightarrow> bool" where |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
62045
diff
changeset
|
67 |
\<comment> \<open>confluence; Lambda/Commutation treats this more abstractly\<close> |
71591 | 68 |
"diamond r \<equiv> \<forall>x y. r x y \<longrightarrow> |
69 |
(\<forall>y'. r x y' \<longrightarrow> |
|
70 |
(\<exists>z. r y z \<and> r y' z))" |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
71 |
|
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
72 |
|
60530 | 73 |
subsection \<open>Reflexive/Transitive closure preserves Church-Rosser property\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
74 |
|
71591 | 75 |
text\<open>Remark: So does the Transitive closure, with a similar proof\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
76 |
|
60530 | 77 |
text\<open>Strip lemma. |
78 |
The induction hypothesis covers all but the last diamond of the strip.\<close> |
|
71591 | 79 |
lemma strip_lemma [rule_format]: |
80 |
assumes "diamond r" and r: "r\<^sup>*\<^sup>* x y" "r x y'" |
|
81 |
shows "\<exists>z. r\<^sup>*\<^sup>* y' z \<and> r y z" |
|
82 |
using r |
|
83 |
proof (induction rule: rtranclp_induct) |
|
84 |
case base |
|
85 |
then show ?case |
|
86 |
by blast |
|
87 |
next |
|
88 |
case (step y z) |
|
89 |
then show ?case |
|
90 |
using \<open>diamond r\<close> unfolding diamond_def |
|
91 |
by (metis rtranclp.rtrancl_into_rtrancl) |
|
92 |
qed |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
93 |
|
71591 | 94 |
proposition diamond_rtrancl: |
95 |
assumes "diamond r" |
|
96 |
shows "diamond(r\<^sup>*\<^sup>*)" |
|
97 |
unfolding diamond_def |
|
98 |
proof (intro strip) |
|
99 |
fix x y y' |
|
100 |
assume "r\<^sup>*\<^sup>* x y" "r\<^sup>*\<^sup>* x y'" |
|
101 |
then show "\<exists>z. r\<^sup>*\<^sup>* y z \<and> r\<^sup>*\<^sup>* y' z" |
|
102 |
proof (induction rule: rtranclp_induct) |
|
103 |
case base |
|
104 |
then show ?case |
|
105 |
by blast |
|
106 |
next |
|
107 |
case (step y z) |
|
108 |
then show ?case |
|
109 |
by (meson assms strip_lemma rtranclp.rtrancl_into_rtrancl) |
|
110 |
qed |
|
111 |
qed |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
112 |
|
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
113 |
|
60530 | 114 |
subsection \<open>Non-contraction results\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
115 |
|
60530 | 116 |
text \<open>Derive a case for each combinator constructor.\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
117 |
|
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
118 |
inductive_cases |
71591 | 119 |
K_contractE [elim!]: "K \<rightarrow>\<^sup>1 r" |
61993 | 120 |
and S_contractE [elim!]: "S \<rightarrow>\<^sup>1 r" |
121 |
and Ap_contractE [elim!]: "p\<bullet>q \<rightarrow>\<^sup>1 r" |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
122 |
|
71591 | 123 |
declare contract1.K [intro!] contract1.S [intro!] |
124 |
declare contract1.Ap1 [intro] contract1.Ap2 [intro] |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
125 |
|
71591 | 126 |
lemma I_contract_E [iff]: "\<not> I \<rightarrow>\<^sup>1 z" |
127 |
unfolding I_def by blast |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
128 |
|
71591 | 129 |
lemma K1_contractD [elim!]: "K\<bullet>x \<rightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<rightarrow>\<^sup>1 x')" |
130 |
by blast |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
131 |
|
61993 | 132 |
lemma Ap_reduce1 [intro]: "x \<rightarrow> y \<Longrightarrow> x\<bullet>z \<rightarrow> y\<bullet>z" |
71591 | 133 |
by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
134 |
|
61993 | 135 |
lemma Ap_reduce2 [intro]: "x \<rightarrow> y \<Longrightarrow> z\<bullet>x \<rightarrow> z\<bullet>y" |
71591 | 136 |
by (induction rule: rtranclp_induct; blast intro: rtranclp_trans) |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
137 |
|
69597 | 138 |
text \<open>Counterexample to the diamond property for \<^term>\<open>x \<rightarrow>\<^sup>1 y\<close>\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
139 |
|
71591 | 140 |
lemma not_diamond_contract: "\<not> diamond(contract1)" |
141 |
unfolding diamond_def by (metis S_contractE contract1.K) |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
142 |
|
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
143 |
|
60530 | 144 |
subsection \<open>Results about Parallel Contraction\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
145 |
|
60530 | 146 |
text \<open>Derive a case for each combinator constructor.\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
147 |
|
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
148 |
inductive_cases |
61993 | 149 |
K_parcontractE [elim!]: "K \<Rrightarrow>\<^sup>1 r" |
150 |
and S_parcontractE [elim!]: "S \<Rrightarrow>\<^sup>1 r" |
|
151 |
and Ap_parcontractE [elim!]: "p\<bullet>q \<Rrightarrow>\<^sup>1 r" |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
152 |
|
71591 | 153 |
declare parcontract1.intros [intro] |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
154 |
|
60530 | 155 |
subsection \<open>Basic properties of parallel contraction\<close> |
71591 | 156 |
text\<open>The rules below are not essential but make proofs much faster\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
157 |
|
71591 | 158 |
lemma K1_parcontractD [dest!]: "K\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = K\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')" |
159 |
by blast |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
160 |
|
71591 | 161 |
lemma S1_parcontractD [dest!]: "S\<bullet>x \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x'. z = S\<bullet>x' \<and> x \<Rrightarrow>\<^sup>1 x')" |
162 |
by blast |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
163 |
|
71591 | 164 |
lemma S2_parcontractD [dest!]: "S\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 z \<Longrightarrow> (\<exists>x' y'. z = S\<bullet>x'\<bullet>y' \<and> x \<Rrightarrow>\<^sup>1 x' \<and> y \<Rrightarrow>\<^sup>1 y')" |
165 |
by blast |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
166 |
|
60530 | 167 |
text\<open>Church-Rosser property for parallel contraction\<close> |
71591 | 168 |
proposition diamond_parcontract: "diamond parcontract1" |
169 |
proof - |
|
170 |
have "(\<exists>z. w \<Rrightarrow>\<^sup>1 z \<and> y' \<Rrightarrow>\<^sup>1 z)" if "y \<Rrightarrow>\<^sup>1 w" "y \<Rrightarrow>\<^sup>1 y'" for w y y' |
|
171 |
using that by (induction arbitrary: y' rule: parcontract1.induct) fast+ |
|
172 |
then show ?thesis |
|
173 |
by (auto simp: diamond_def) |
|
174 |
qed |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
175 |
|
71591 | 176 |
subsection \<open>Equivalence of \<^prop>\<open>p \<rightarrow> q\<close> and \<^prop>\<open>p \<Rrightarrow> q\<close>.\<close> |
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
177 |
|
71591 | 178 |
lemma contract_imp_parcontract: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> x \<Rrightarrow>\<^sup>1 y" |
179 |
by (induction rule: contract1.induct; blast) |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
180 |
|
60530 | 181 |
text\<open>Reductions: simply throw together reflexivity, transitivity and |
182 |
the one-step reductions\<close> |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
183 |
|
71591 | 184 |
proposition reduce_I: "I\<bullet>x \<rightarrow> x" |
185 |
unfolding I_def |
|
186 |
by (meson contract1.K contract1.S r_into_rtranclp rtranclp.rtrancl_into_rtrancl) |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
187 |
|
71591 | 188 |
lemma parcontract_imp_reduce: "x \<Rrightarrow>\<^sup>1 y \<Longrightarrow> x \<rightarrow> y" |
189 |
proof (induction rule: parcontract1.induct) |
|
190 |
case (Ap x y z w) |
|
191 |
then show ?case |
|
192 |
by (meson Ap_reduce1 Ap_reduce2 rtranclp_trans) |
|
193 |
qed auto |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
194 |
|
71591 | 195 |
lemma reduce_eq_parreduce: "x \<rightarrow> y \<longleftrightarrow> x \<Rrightarrow> y" |
196 |
by (metis contract_imp_parcontract parcontract_imp_reduce predicate2I rtranclp_subset) |
|
13075
d3e1d554cd6d
conversion of some HOL/Induct proof scripts to Isar
paulson
parents:
11359
diff
changeset
|
197 |
|
71591 | 198 |
theorem diamond_reduce: "diamond(contract)" |
199 |
using diamond_parcontract diamond_rtrancl reduce_eq_parreduce by presburger |
|
3120
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
200 |
|
c58423c20740
New directory to contain examples of (co)inductive definitions
paulson
parents:
diff
changeset
|
201 |
end |