| author | webertj | 
| Tue, 01 Aug 2006 14:58:43 +0200 | |
| changeset 20276 | d94dc40673b1 | 
| parent 19669 | 95ac857276e1 | 
| child 21404 | eb85850d3eb7 | 
| 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 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 20 | consts | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 21 |   acc :: "('a \<times> 'a) set => 'a set"
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 22 | inductive "acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 23 | intros | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 24 | accI: "(!!y. (y, x) \<in> r ==> y \<in> acc r) ==> x \<in> acc r" | 
| 
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 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 27 |   termi :: "('a \<times> 'a) set => 'a set"
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 28 | "termi r == acc (r\<inverse>)" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 29 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 30 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 31 | subsection {* Induction rules *}
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 32 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 33 | theorem acc_induct: | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 34 | assumes major: "a \<in> acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 35 | assumes hyp: "!!x. x \<in> acc r ==> \<forall>y. (y, x) \<in> r --> P y ==> P x" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 36 | shows "P a" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 37 | apply (rule major [THEN acc.induct]) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 38 | apply (rule hyp) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 39 | apply (rule accI) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 40 | apply fast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 41 | apply fast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 42 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 43 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 44 | theorems acc_induct_rule = acc_induct [rule_format, induct set: acc] | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 45 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 46 | theorem acc_downward: "b \<in> acc r ==> (a, b) \<in> r ==> a \<in> acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 47 | apply (erule acc.elims) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 48 | apply fast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 49 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 50 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 51 | lemma acc_downwards_aux: "(b, a) \<in> r\<^sup>* ==> a \<in> acc r --> b \<in> acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 52 | apply (erule rtrancl_induct) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 53 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 54 | apply (blast dest: acc_downward) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 55 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 56 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 57 | theorem acc_downwards: "a \<in> acc r ==> (b, a) \<in> r\<^sup>* ==> b \<in> acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 58 | apply (blast dest: acc_downwards_aux) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 59 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 60 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 61 | theorem acc_wfI: "\<forall>x. x \<in> acc r ==> wf r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 62 | apply (rule wfUNIVI) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 63 | apply (induct_tac P x rule: acc_induct) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 64 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 65 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 66 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 67 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 68 | theorem acc_wfD: "wf r ==> x \<in> acc r" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 69 | apply (erule wf_induct) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 70 | apply (rule accI) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 71 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 72 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 73 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 74 | theorem wf_acc_iff: "wf r = (\<forall>x. x \<in> acc r)" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 75 | apply (blast intro: acc_wfI dest: acc_wfD) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 76 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 77 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 78 | |
| 19669 | 79 | text {* Smaller relations have bigger accessible parts: *}
 | 
| 80 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 81 | lemma acc_subset: | 
| 19669 | 82 | assumes sub: "R1 \<subseteq> R2" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 83 | shows "acc R2 \<subseteq> acc R1" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 84 | proof | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 85 | fix x assume "x \<in> acc R2" | 
| 19669 | 86 | then show "x \<in> acc R1" | 
| 87 | proof (induct x) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 88 | fix x | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 89 | assume ih: "\<And>y. (y, x) \<in> R2 \<Longrightarrow> y \<in> acc R1" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 90 | with sub show "x \<in> acc R1" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 91 | by (blast intro:accI) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 92 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 93 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 94 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 95 | |
| 19669 | 96 | text {* This is a generalized induction theorem that works on
 | 
| 97 | subsets of the accessible part. *} | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 98 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 99 | lemma acc_subset_induct: | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 100 | assumes subset: "D \<subseteq> acc R" | 
| 19669 | 101 | and dcl: "\<And>x z. \<lbrakk>x \<in> D; (z, x)\<in>R\<rbrakk> \<Longrightarrow> z \<in> D" | 
| 102 | and "x \<in> D" | |
| 103 | and istep: "\<And>x. \<lbrakk>x \<in> D; (\<And>z. (z, x)\<in>R \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x" | |
| 104 | shows "P x" | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 105 | proof - | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 106 | from `x \<in> D` and subset | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 107 | have "x \<in> acc R" .. | 
| 19669 | 108 | then show "P x" using `x \<in> D` | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 109 | proof (induct x) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 110 | fix x | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 111 | assume "x \<in> D" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 112 | and "\<And>y. (y, x) \<in> R \<Longrightarrow> y \<in> D \<Longrightarrow> P y" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 113 | 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 | 114 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 115 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 116 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 117 | end |