| author | berghofe | 
| Wed, 11 Jul 2007 11:04:39 +0200 | |
| changeset 23740 | d7f18c837ce7 | 
| parent 23735 | afc12f93f64f | 
| child 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 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 34 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 35 | subsection {* Induction rules *}
 | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 36 | |
| 23735 | 37 | theorem accp_induct: | 
| 38 | assumes major: "accp r a" | |
| 39 | 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 | 40 | shows "P a" | 
| 23735 | 41 | apply (rule major [THEN accp.induct]) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 42 | apply (rule hyp) | 
| 23735 | 43 | apply (rule accp.accI) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 44 | apply fast | 
| 
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 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 47 | |
| 23735 | 48 | 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 | 49 | |
| 23735 | 50 | theorem accp_downward: "accp r b ==> r a b ==> accp r a" | 
| 51 | apply (erule accp.cases) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 52 | apply fast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 53 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 54 | |
| 23735 | 55 | lemma not_accp_down: | 
| 56 | assumes na: "\<not> accp R x" | |
| 57 | obtains z where "R z x" and "\<not> accp R z" | |
| 22262 | 58 | proof - | 
| 23735 | 59 | assume a: "\<And>z. \<lbrakk>R z x; \<not> accp R z\<rbrakk> \<Longrightarrow> thesis" | 
| 22262 | 60 | |
| 61 | show thesis | |
| 23735 | 62 | proof (cases "\<forall>z. R z x \<longrightarrow> accp R z") | 
| 22262 | 63 | case True | 
| 23735 | 64 | hence "\<And>z. R z x \<Longrightarrow> accp R z" by auto | 
| 65 | hence "accp R x" | |
| 66 | by (rule accp.accI) | |
| 22262 | 67 | with na show thesis .. | 
| 68 | next | |
| 23735 | 69 | case False then obtain z where "R z x" and "\<not> accp R z" | 
| 22262 | 70 | by auto | 
| 71 | with a show thesis . | |
| 72 | qed | |
| 73 | qed | |
| 74 | ||
| 23735 | 75 | lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a ==> accp r a --> accp r b" | 
| 76 | apply (erule rtranclp_induct) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 77 | apply blast | 
| 23735 | 78 | apply (blast dest: accp_downward) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 79 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 80 | |
| 23735 | 81 | theorem accp_downwards: "accp r a ==> r\<^sup>*\<^sup>* b a ==> accp r b" | 
| 82 | apply (blast dest: accp_downwards_aux) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 83 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 84 | |
| 23735 | 85 | theorem accp_wfPI: "\<forall>x. accp r x ==> wfP r" | 
| 22262 | 86 | apply (rule wfPUNIVI) | 
| 23735 | 87 | 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 | 88 | apply blast | 
| 
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 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 91 | |
| 23735 | 92 | theorem accp_wfPD: "wfP r ==> accp r x" | 
| 22262 | 93 | apply (erule wfP_induct_rule) | 
| 23735 | 94 | apply (rule accp.accI) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 95 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 96 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 97 | |
| 23735 | 98 | theorem wfP_accp_iff: "wfP r = (\<forall>x. accp r x)" | 
| 99 | 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 | 100 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 101 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 102 | |
| 19669 | 103 | text {* Smaller relations have bigger accessible parts: *}
 | 
| 104 | ||
| 23735 | 105 | lemma accp_subset: | 
| 22262 | 106 | assumes sub: "R1 \<le> R2" | 
| 23735 | 107 | shows "accp R2 \<le> accp R1" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 108 | proof | 
| 23735 | 109 | fix x assume "accp R2 x" | 
| 110 | then show "accp R1 x" | |
| 19669 | 111 | proof (induct x) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 112 | fix x | 
| 23735 | 113 | assume ih: "\<And>y. R2 y x \<Longrightarrow> accp R1 y" | 
| 114 | with sub show "accp R1 x" | |
| 115 | by (blast intro: accp.accI) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 116 | qed | 
| 
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 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 119 | |
| 19669 | 120 | text {* This is a generalized induction theorem that works on
 | 
| 121 | subsets of the accessible part. *} | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 122 | |
| 23735 | 123 | lemma accp_subset_induct: | 
| 124 | assumes subset: "D \<le> accp R" | |
| 22262 | 125 | and dcl: "\<And>x z. \<lbrakk>D x; R z x\<rbrakk> \<Longrightarrow> D z" | 
| 126 | and "D x" | |
| 127 | and istep: "\<And>x. \<lbrakk>D x; (\<And>z. R z x \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x" | |
| 19669 | 128 | shows "P x" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 129 | proof - | 
| 22262 | 130 | from subset and `D x` | 
| 23735 | 131 | have "accp R x" .. | 
| 22262 | 132 | 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 | 133 | proof (induct x) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 134 | fix x | 
| 22262 | 135 | assume "D x" | 
| 136 | 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 | 137 | 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 | 138 | qed | 
| 
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 | |
| 23735 | 141 | |
| 142 | text {* Set versions of the above theorems *}
 | |
| 143 | ||
| 144 | lemmas acc_induct = accp_induct [to_set] | |
| 145 | ||
| 146 | lemmas acc_induct_rule = acc_induct [rule_format, induct set: acc] | |
| 147 | ||
| 148 | lemmas acc_downward = accp_downward [to_set] | |
| 149 | ||
| 150 | lemmas not_acc_down = not_accp_down [to_set] | |
| 151 | ||
| 152 | lemmas acc_downwards_aux = accp_downwards_aux [to_set] | |
| 153 | ||
| 154 | lemmas acc_downwards = accp_downwards [to_set] | |
| 155 | ||
| 156 | lemmas acc_wfI = accp_wfPI [to_set] | |
| 157 | ||
| 158 | lemmas acc_wfD = accp_wfPD [to_set] | |
| 159 | ||
| 160 | lemmas wf_acc_iff = wfP_accp_iff [to_set] | |
| 161 | ||
| 162 | lemmas acc_subset = accp_subset [to_set] | |
| 163 | ||
| 164 | lemmas acc_subset_induct = accp_subset_induct [to_set] | |
| 165 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 166 | end |