| author | urbanc | 
| Thu, 22 Mar 2007 11:17:32 +0100 | |
| changeset 22495 | c54748fd1f43 | 
| parent 22262 | 96ba62dff413 | 
| child 23735 | afc12f93f64f | 
| 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 | |
| 22262 | 20 | inductive2 | 
| 21 |   acc :: "('a => 'a => bool) => 'a => bool"
 | |
| 22 | for r :: "'a => 'a => bool" | |
| 23 | where | |
| 24 | accI: "(!!y. r y x ==> acc r y) ==> acc r x" | |
| 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 | 
| 22262 | 27 |   termi :: "('a => 'a => bool) => 'a => bool" where
 | 
| 28 | "termi r == acc (r\<inverse>\<inverse>)" | |
| 19564 
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: | 
| 22262 | 34 | assumes major: "acc r a" | 
| 35 | assumes hyp: "!!x. acc 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 | 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 | |
| 22262 | 46 | theorem acc_downward: "acc r b ==> r a b ==> acc r a" | 
| 47 | apply (erule acc.cases) | |
| 19564 
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 | |
| 22262 | 51 | lemma not_acc_down: | 
| 52 | assumes na: "\<not> acc R x" | |
| 53 | obtains z where "R z x" and "\<not> acc R z" | |
| 54 | proof - | |
| 55 | assume a: "\<And>z. \<lbrakk>R z x; \<not> acc R z\<rbrakk> \<Longrightarrow> thesis" | |
| 56 | ||
| 57 | show thesis | |
| 58 | proof (cases "\<forall>z. R z x \<longrightarrow> acc R z") | |
| 59 | case True | |
| 60 | hence "\<And>z. R z x \<Longrightarrow> acc R z" by auto | |
| 61 | hence "acc R x" | |
| 62 | by (rule accI) | |
| 63 | with na show thesis .. | |
| 64 | next | |
| 65 | case False then obtain z where "R z x" and "\<not> acc R z" | |
| 66 | by auto | |
| 67 | with a show thesis . | |
| 68 | qed | |
| 69 | qed | |
| 70 | ||
| 71 | lemma acc_downwards_aux: "r\<^sup>*\<^sup>* b a ==> acc r a --> acc r b" | |
| 72 | apply (erule rtrancl_induct') | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 73 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 74 | apply (blast dest: acc_downward) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 75 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 76 | |
| 22262 | 77 | theorem acc_downwards: "acc r a ==> r\<^sup>*\<^sup>* b a ==> acc r b" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 78 | apply (blast dest: acc_downwards_aux) | 
| 
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 | |
| 22262 | 81 | theorem acc_wfI: "\<forall>x. acc r x ==> wfP r" | 
| 82 | apply (rule wfPUNIVI) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 83 | 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 | 84 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 85 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 86 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 87 | |
| 22262 | 88 | theorem acc_wfD: "wfP r ==> acc r x" | 
| 89 | apply (erule wfP_induct_rule) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 90 | apply (rule accI) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 91 | apply blast | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 92 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 93 | |
| 22262 | 94 | theorem wf_acc_iff: "wfP r = (\<forall>x. acc r x)" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 95 | 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 | 96 | done | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 97 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 98 | |
| 19669 | 99 | text {* Smaller relations have bigger accessible parts: *}
 | 
| 100 | ||
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 101 | lemma acc_subset: | 
| 22262 | 102 | assumes sub: "R1 \<le> R2" | 
| 103 | shows "acc R2 \<le> acc R1" | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 104 | proof | 
| 22262 | 105 | fix x assume "acc R2 x" | 
| 106 | then show "acc R1 x" | |
| 19669 | 107 | proof (induct x) | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 108 | fix x | 
| 22262 | 109 | assume ih: "\<And>y. R2 y x \<Longrightarrow> acc R1 y" | 
| 110 | with sub show "acc R1 x" | |
| 111 | by (blast intro: accI) | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 112 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 113 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 114 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 115 | |
| 19669 | 116 | text {* This is a generalized induction theorem that works on
 | 
| 117 | subsets of the accessible part. *} | |
| 19564 
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 | lemma acc_subset_induct: | 
| 22262 | 120 | assumes subset: "D \<le> acc R" | 
| 121 | and dcl: "\<And>x z. \<lbrakk>D x; R z x\<rbrakk> \<Longrightarrow> D z" | |
| 122 | and "D x" | |
| 123 | and istep: "\<And>x. \<lbrakk>D x; (\<And>z. R z x \<Longrightarrow> P z)\<rbrakk> \<Longrightarrow> P x" | |
| 19669 | 124 | shows "P x" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 125 | proof - | 
| 22262 | 126 | from subset and `D x` | 
| 127 | have "acc R x" .. | |
| 128 | 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 | 129 | proof (induct x) | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 130 | fix x | 
| 22262 | 131 | assume "D x" | 
| 132 | 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 | 133 | 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 | 134 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 135 | qed | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 136 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 137 | end |