| author | wenzelm | 
| Fri, 31 Aug 2007 23:17:25 +0200 | |
| changeset 24505 | 9e6d91f8bb73 | 
| parent 23818 | cfe8d4bf749a | 
| permissions | -rw-r--r-- | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Accessible_Part.thy  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
4  | 
Copyright 1994 University of Cambridge  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
7  | 
header {* The accessible part of a relation *}
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
9  | 
theory Accessible_Part  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
10  | 
imports Wellfounded_Recursion  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
11  | 
begin  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
12  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
13  | 
subsection {* Inductive definition *}
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
15  | 
text {*
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
16  | 
 Inductive definition of the accessible part @{term "acc r"} of a
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
17  | 
 relation; see also \cite{paulin-tlca}.
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
18  | 
*}  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
19  | 
|
| 23735 | 20  | 
inductive_set  | 
21  | 
  acc :: "('a * 'a) set => 'a set"
 | 
|
22  | 
  for r :: "('a * 'a) set"
 | 
|
| 22262 | 23  | 
where  | 
| 23735 | 24  | 
accI: "(!!y. (y, x) : r ==> y : acc r) ==> x : acc r"  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
26  | 
abbreviation  | 
| 23735 | 27  | 
  termip :: "('a => 'a => bool) => 'a => bool" where
 | 
28  | 
"termip r == accp (r\<inverse>\<inverse>)"  | 
|
29  | 
||
30  | 
abbreviation  | 
|
31  | 
  termi :: "('a * 'a) set => 'a set" where
 | 
|
32  | 
"termi r == acc (r\<inverse>)"  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
33  | 
|
| 23818 | 34  | 
lemmas accpI = accp.accI  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
36  | 
subsection {* Induction rules *}
 | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
37  | 
|
| 23735 | 38  | 
theorem accp_induct:  | 
39  | 
assumes major: "accp r a"  | 
|
40  | 
assumes hyp: "!!x. accp r x ==> \<forall>y. r y x --> P y ==> P x"  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
41  | 
shows "P a"  | 
| 23735 | 42  | 
apply (rule major [THEN accp.induct])  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
43  | 
apply (rule hyp)  | 
| 23735 | 44  | 
apply (rule accp.accI)  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
45  | 
apply fast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
46  | 
apply fast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
47  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
48  | 
|
| 23735 | 49  | 
theorems accp_induct_rule = accp_induct [rule_format, induct set: accp]  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
50  | 
|
| 23735 | 51  | 
theorem accp_downward: "accp r b ==> r a b ==> accp r a"  | 
52  | 
apply (erule accp.cases)  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
53  | 
apply fast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
54  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
55  | 
|
| 23735 | 56  | 
lemma not_accp_down:  | 
57  | 
assumes na: "\<not> accp R x"  | 
|
58  | 
obtains z where "R z x" and "\<not> accp R z"  | 
|
| 22262 | 59  | 
proof -  | 
| 23735 | 60  | 
assume a: "\<And>z. \<lbrakk>R z x; \<not> accp R z\<rbrakk> \<Longrightarrow> thesis"  | 
| 22262 | 61  | 
|
62  | 
show thesis  | 
|
| 23735 | 63  | 
proof (cases "\<forall>z. R z x \<longrightarrow> accp R z")  | 
| 22262 | 64  | 
case True  | 
| 23735 | 65  | 
hence "\<And>z. R z x \<Longrightarrow> accp R z" by auto  | 
66  | 
hence "accp R x"  | 
|
67  | 
by (rule accp.accI)  | 
|
| 22262 | 68  | 
with na show thesis ..  | 
69  | 
next  | 
|
| 23735 | 70  | 
case False then obtain z where "R z x" and "\<not> accp R z"  | 
| 22262 | 71  | 
by auto  | 
72  | 
with a show thesis .  | 
|
73  | 
qed  | 
|
74  | 
qed  | 
|
75  | 
||
| 23735 | 76  | 
lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a ==> accp r a --> accp r b"  | 
77  | 
apply (erule rtranclp_induct)  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
78  | 
apply blast  | 
| 23735 | 79  | 
apply (blast dest: accp_downward)  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
80  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
81  | 
|
| 23735 | 82  | 
theorem accp_downwards: "accp r a ==> r\<^sup>*\<^sup>* b a ==> accp r b"  | 
83  | 
apply (blast dest: accp_downwards_aux)  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
84  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
85  | 
|
| 23735 | 86  | 
theorem accp_wfPI: "\<forall>x. accp r x ==> wfP r"  | 
| 22262 | 87  | 
apply (rule wfPUNIVI)  | 
| 23735 | 88  | 
apply (induct_tac P x rule: accp_induct)  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
89  | 
apply blast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
90  | 
apply blast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
91  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
92  | 
|
| 23735 | 93  | 
theorem accp_wfPD: "wfP r ==> accp r x"  | 
| 22262 | 94  | 
apply (erule wfP_induct_rule)  | 
| 23735 | 95  | 
apply (rule accp.accI)  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
96  | 
apply blast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
97  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
98  | 
|
| 23735 | 99  | 
theorem wfP_accp_iff: "wfP r = (\<forall>x. accp r x)"  | 
100  | 
apply (blast intro: accp_wfPI dest: accp_wfPD)  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
101  | 
done  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
103  | 
|
| 19669 | 104  | 
text {* Smaller relations have bigger accessible parts: *}
 | 
105  | 
||
| 23735 | 106  | 
lemma accp_subset:  | 
| 22262 | 107  | 
assumes sub: "R1 \<le> R2"  | 
| 23735 | 108  | 
shows "accp R2 \<le> accp R1"  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
109  | 
proof  | 
| 23735 | 110  | 
fix x assume "accp R2 x"  | 
111  | 
then show "accp R1 x"  | 
|
| 19669 | 112  | 
proof (induct x)  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
113  | 
fix x  | 
| 23735 | 114  | 
assume ih: "\<And>y. R2 y x \<Longrightarrow> accp R1 y"  | 
115  | 
with sub show "accp R1 x"  | 
|
116  | 
by (blast intro: accp.accI)  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
117  | 
qed  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
118  | 
qed  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
120  | 
|
| 19669 | 121  | 
text {* This is a generalized induction theorem that works on
 | 
122  | 
subsets of the accessible part. *}  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
123  | 
|
| 23735 | 124  | 
lemma accp_subset_induct:  | 
125  | 
assumes subset: "D \<le> accp R"  | 
|
| 22262 | 126  | 
and dcl: "\<And>x z. \<lbrakk>D x; R z x\<rbrakk> \<Longrightarrow> D z"  | 
127  | 
and "D x"  | 
|
128  | 
and istep: "\<And>x. \<lbrakk>D x; (\<And>z. R z x \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x"  | 
|
| 19669 | 129  | 
shows "P x"  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
130  | 
proof -  | 
| 22262 | 131  | 
from subset and `D x`  | 
| 23735 | 132  | 
have "accp R x" ..  | 
| 22262 | 133  | 
then show "P x" using `D x`  | 
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
134  | 
proof (induct x)  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
135  | 
fix x  | 
| 22262 | 136  | 
assume "D x"  | 
137  | 
and "\<And>y. R y x \<Longrightarrow> D y \<Longrightarrow> P y"  | 
|
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
138  | 
with dcl and istep show "P x" by blast  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
139  | 
qed  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
140  | 
qed  | 
| 
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
141  | 
|
| 23735 | 142  | 
|
143  | 
text {* Set versions of the above theorems *}
 | 
|
144  | 
||
145  | 
lemmas acc_induct = accp_induct [to_set]  | 
|
146  | 
||
147  | 
lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc]  | 
|
148  | 
||
149  | 
lemmas acc_downward = accp_downward [to_set]  | 
|
150  | 
||
151  | 
lemmas not_acc_down = not_accp_down [to_set]  | 
|
152  | 
||
153  | 
lemmas acc_downwards_aux = accp_downwards_aux [to_set]  | 
|
154  | 
||
155  | 
lemmas acc_downwards = accp_downwards [to_set]  | 
|
156  | 
||
157  | 
lemmas acc_wfI = accp_wfPI [to_set]  | 
|
158  | 
||
159  | 
lemmas acc_wfD = accp_wfPD [to_set]  | 
|
160  | 
||
161  | 
lemmas wf_acc_iff = wfP_accp_iff [to_set]  | 
|
162  | 
||
163  | 
lemmas acc_subset = accp_subset [to_set]  | 
|
164  | 
||
165  | 
lemmas acc_subset_induct = accp_subset_induct [to_set]  | 
|
166  | 
||
| 
19564
 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 
krauss 
parents:  
diff
changeset
 | 
167  | 
end  |