| author | nipkow | 
| Thu, 01 Apr 2010 09:31:58 +0200 | |
| changeset 36070 | d80e5d3c8fe1 | 
| parent 35416 | d8d7d1b785af | 
| child 37767 | a2b7a20d6ea3 | 
| permissions | -rw-r--r-- | 
| 7701 | 1 | (* Title: HOL/Recdef.thy | 
| 10773 | 2 | Author: Konrad Slind and Markus Wenzel, TU Muenchen | 
| 12023 | 3 | *) | 
| 5123 | 4 | |
| 12023 | 5 | header {* TFL: recursive function definitions *}
 | 
| 7701 | 6 | |
| 15131 | 7 | theory Recdef | 
| 29654 
24e73987bfe2
Plain, Main form meeting points in import hierarchy
 haftmann parents: 
26748diff
changeset | 8 | imports FunDef Plain | 
| 16417 | 9 | uses | 
| 23150 | 10 |   ("Tools/TFL/casesplit.ML")
 | 
| 11 |   ("Tools/TFL/utils.ML")
 | |
| 12 |   ("Tools/TFL/usyntax.ML")
 | |
| 13 |   ("Tools/TFL/dcterm.ML")
 | |
| 14 |   ("Tools/TFL/thms.ML")
 | |
| 15 |   ("Tools/TFL/rules.ML")
 | |
| 16 |   ("Tools/TFL/thry.ML")
 | |
| 17 |   ("Tools/TFL/tfl.ML")
 | |
| 18 |   ("Tools/TFL/post.ML")
 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
29654diff
changeset | 19 |   ("Tools/recdef.ML")
 | 
| 15131 | 20 | begin | 
| 10773 | 21 | |
| 32462 | 22 | inductive | 
| 23 |   wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool"
 | |
| 24 |   for R :: "('a * 'a) set"
 | |
| 25 |   and F :: "('a => 'b) => 'a => 'b"
 | |
| 26 | where | |
| 27 | wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==> | |
| 28 | wfrec_rel R F x (F g x)" | |
| 29 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 30 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 31 |   cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" where
 | 
| 32462 | 32 | "cut f r x == (%y. if (y,x):r then f y else undefined)" | 
| 33 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 34 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 35 |   adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" where
 | 
| 32462 | 36 | "adm_wf R F == ALL f g x. | 
| 37 | (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x" | |
| 38 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 39 | definition | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32462diff
changeset | 40 |   wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" where
 | 
| 32462 | 41 | [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y" | 
| 42 | ||
| 43 | subsection{*Well-Founded Recursion*}
 | |
| 44 | ||
| 45 | text{*cut*}
 | |
| 46 | ||
| 47 | lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))" | |
| 48 | by (simp add: expand_fun_eq cut_def) | |
| 49 | ||
| 50 | lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)" | |
| 51 | by (simp add: cut_def) | |
| 52 | ||
| 53 | text{*Inductive characterization of wfrec combinator; for details see:
 | |
| 54 | John Harrison, "Inductive definitions: automation and application"*} | |
| 55 | ||
| 56 | lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y" | |
| 57 | apply (simp add: adm_wf_def) | |
| 58 | apply (erule_tac a=x in wf_induct) | |
| 59 | apply (rule ex1I) | |
| 60 | apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI) | |
| 61 | apply (fast dest!: theI') | |
| 62 | apply (erule wfrec_rel.cases, simp) | |
| 63 | apply (erule allE, erule allE, erule allE, erule mp) | |
| 64 | apply (fast intro: the_equality [symmetric]) | |
| 65 | done | |
| 66 | ||
| 67 | lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)" | |
| 68 | apply (simp add: adm_wf_def) | |
| 69 | apply (intro strip) | |
| 70 | apply (rule cuts_eq [THEN iffD2, THEN subst], assumption) | |
| 71 | apply (rule refl) | |
| 72 | done | |
| 73 | ||
| 74 | lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a" | |
| 75 | apply (simp add: wfrec_def) | |
| 76 | apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption) | |
| 77 | apply (rule wfrec_rel.wfrecI) | |
| 78 | apply (intro strip) | |
| 79 | apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) | |
| 80 | done | |
| 81 | ||
| 82 | ||
| 26748 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 83 | text{** This form avoids giant explosions in proofs.  NOTE USE OF ==*}
 | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 84 | lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a" | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 85 | apply auto | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 86 | apply (blast intro: wfrec) | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 87 | done | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 88 | |
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 89 | |
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 90 | lemma tfl_wf_induct: "ALL R. wf R --> | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 91 | (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))" | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 92 | apply clarify | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 93 | apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast) | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 94 | done | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 95 | |
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 96 | lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)" | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 97 | apply clarify | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 98 | apply (rule cut_apply, assumption) | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 99 | done | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 100 | |
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 101 | lemma tfl_wfrec: | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 102 | "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)" | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 103 | apply clarify | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 104 | apply (erule wfrec) | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 105 | done | 
| 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 krauss parents: 
23742diff
changeset | 106 | |
| 10773 | 107 | lemma tfl_eq_True: "(x = True) --> x" | 
| 108 | by blast | |
| 109 | ||
| 110 | lemma tfl_rev_eq_mp: "(x = y) --> y --> x"; | |
| 111 | by blast | |
| 112 | ||
| 113 | lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)" | |
| 114 | by blast | |
| 6438 | 115 | |
| 10773 | 116 | lemma tfl_P_imp_P_iff_True: "P ==> P = True" | 
| 117 | by blast | |
| 118 | ||
| 119 | lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)" | |
| 120 | by blast | |
| 121 | ||
| 12023 | 122 | lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)" | 
| 10773 | 123 | by simp | 
| 124 | ||
| 12023 | 125 | lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R" | 
| 10773 | 126 | by blast | 
| 127 | ||
| 12023 | 128 | lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q" | 
| 10773 | 129 | by blast | 
| 130 | ||
| 23150 | 131 | use "Tools/TFL/casesplit.ML" | 
| 132 | use "Tools/TFL/utils.ML" | |
| 133 | use "Tools/TFL/usyntax.ML" | |
| 134 | use "Tools/TFL/dcterm.ML" | |
| 135 | use "Tools/TFL/thms.ML" | |
| 136 | use "Tools/TFL/rules.ML" | |
| 137 | use "Tools/TFL/thry.ML" | |
| 138 | use "Tools/TFL/tfl.ML" | |
| 139 | use "Tools/TFL/post.ML" | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
29654diff
changeset | 140 | use "Tools/recdef.ML" | 
| 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
29654diff
changeset | 141 | setup Recdef.setup | 
| 6438 | 142 | |
| 32244 | 143 | text {*Wellfoundedness of @{text same_fst}*}
 | 
| 144 | ||
| 145 | definition | |
| 146 |  same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
 | |
| 147 | where | |
| 148 |     "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
 | |
| 149 |    --{*For @{text rec_def} declarations where the first n parameters
 | |
| 150 | stay unchanged in the recursive call. *} | |
| 151 | ||
| 152 | lemma same_fstI [intro!]: | |
| 153 | "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R" | |
| 154 | by (simp add: same_fst_def) | |
| 155 | ||
| 156 | lemma wf_same_fst: | |
| 157 | assumes prem: "(!!x. P x ==> wf(R x))" | |
| 158 | shows "wf(same_fst P R)" | |
| 159 | apply (simp cong del: imp_cong add: wf_def same_fst_def) | |
| 160 | apply (intro strip) | |
| 161 | apply (rename_tac a b) | |
| 162 | apply (case_tac "wf (R a)") | |
| 163 | apply (erule_tac a = b in wf_induct, blast) | |
| 164 | apply (blast intro: prem) | |
| 165 | done | |
| 166 | ||
| 167 | text {*Rule setup*}
 | |
| 168 | ||
| 9855 | 169 | lemmas [recdef_simp] = | 
| 170 | inv_image_def | |
| 171 | measure_def | |
| 172 | lex_prod_def | |
| 11284 | 173 | same_fst_def | 
| 9855 | 174 | less_Suc_eq [THEN iffD2] | 
| 175 | ||
| 23150 | 176 | lemmas [recdef_cong] = | 
| 22622 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22399diff
changeset | 177 | if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong | 
| 9855 | 178 | |
| 179 | lemmas [recdef_wf] = | |
| 180 | wf_trancl | |
| 181 | wf_less_than | |
| 182 | wf_lex_prod | |
| 183 | wf_inv_image | |
| 184 | wf_measure | |
| 185 | wf_pred_nat | |
| 10653 | 186 | wf_same_fst | 
| 9855 | 187 | wf_empty | 
| 188 | ||
| 6438 | 189 | end |