renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
1 (* Title: HOL/Recdef.thy
2 Author: Konrad Slind and Markus Wenzel, TU Muenchen
5 header {* TFL: recursive function definitions *}
10 ("Tools/TFL/casesplit.ML")
11 ("Tools/TFL/utils.ML")
12 ("Tools/TFL/usyntax.ML")
13 ("Tools/TFL/dcterm.ML")
15 ("Tools/TFL/rules.ML")
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"
27 wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
28 wfrec_rel R F x (F g x)"
31 cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
32 "cut f r x == (%y. if (y,x):r then f y else undefined)"
34 adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
35 "adm_wf R F == ALL f g x.
36 (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
38 wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
39 [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
41 subsection{*Well-Founded Recursion*}
45 lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"
46 by (simp add: expand_fun_eq cut_def)
48 lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)"
49 by (simp add: cut_def)
51 text{*Inductive characterization of wfrec combinator; for details see:
52 John Harrison, "Inductive definitions: automation and application"*}
54 lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y"
55 apply (simp add: adm_wf_def)
56 apply (erule_tac a=x in wf_induct)
58 apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI)
59 apply (fast dest!: theI')
60 apply (erule wfrec_rel.cases, simp)
61 apply (erule allE, erule allE, erule allE, erule mp)
62 apply (fast intro: the_equality [symmetric])
65 lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
66 apply (simp add: adm_wf_def)
68 apply (rule cuts_eq [THEN iffD2, THEN subst], assumption)
72 lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"
73 apply (simp add: wfrec_def)
74 apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption)
75 apply (rule wfrec_rel.wfrecI)
77 apply (erule adm_lemma [THEN wfrec_unique, THEN theI'])
81 text{** This form avoids giant explosions in proofs. NOTE USE OF ==*}
82 lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a"
84 apply (blast intro: wfrec)
88 lemma tfl_wf_induct: "ALL R. wf R -->
89 (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"
91 apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast)
94 lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)"
96 apply (rule cut_apply, assumption)
100 "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)"
105 lemma tfl_eq_True: "(x = True) --> x"
108 lemma tfl_rev_eq_mp: "(x = y) --> y --> x";
111 lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)"
114 lemma tfl_P_imp_P_iff_True: "P ==> P = True"
117 lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)"
120 lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)"
123 lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R"
126 lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q"
129 use "Tools/TFL/casesplit.ML"
130 use "Tools/TFL/utils.ML"
131 use "Tools/TFL/usyntax.ML"
132 use "Tools/TFL/dcterm.ML"
133 use "Tools/TFL/thms.ML"
134 use "Tools/TFL/rules.ML"
135 use "Tools/TFL/thry.ML"
136 use "Tools/TFL/tfl.ML"
137 use "Tools/TFL/post.ML"
138 use "Tools/recdef.ML"
141 text {*Wellfoundedness of @{text same_fst}*}
144 same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
146 "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
147 --{*For @{text rec_def} declarations where the first n parameters
148 stay unchanged in the recursive call. *}
150 lemma same_fstI [intro!]:
151 "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
152 by (simp add: same_fst_def)
155 assumes prem: "(!!x. P x ==> wf(R x))"
156 shows "wf(same_fst P R)"
157 apply (simp cong del: imp_cong add: wf_def same_fst_def)
159 apply (rename_tac a b)
160 apply (case_tac "wf (R a)")
161 apply (erule_tac a = b in wf_induct, blast)
162 apply (blast intro: prem)
167 lemmas [recdef_simp] =
172 less_Suc_eq [THEN iffD2]
174 lemmas [recdef_cong] =
175 if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong