wenzelm@7701: (* Title: HOL/Recdef.thy wenzelm@10773: Author: Konrad Slind and Markus Wenzel, TU Muenchen wenzelm@12023: *) wenzelm@5123: wenzelm@12023: header {* TFL: recursive function definitions *} wenzelm@7701: nipkow@15131: theory Recdef haftmann@29654: imports FunDef Plain haftmann@16417: uses wenzelm@23150: ("Tools/TFL/casesplit.ML") wenzelm@23150: ("Tools/TFL/utils.ML") wenzelm@23150: ("Tools/TFL/usyntax.ML") wenzelm@23150: ("Tools/TFL/dcterm.ML") wenzelm@23150: ("Tools/TFL/thms.ML") wenzelm@23150: ("Tools/TFL/rules.ML") wenzelm@23150: ("Tools/TFL/thry.ML") wenzelm@23150: ("Tools/TFL/tfl.ML") wenzelm@23150: ("Tools/TFL/post.ML") haftmann@31723: ("Tools/recdef.ML") nipkow@15131: begin wenzelm@10773: krauss@32462: inductive krauss@32462: wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool" krauss@32462: for R :: "('a * 'a) set" krauss@32462: and F :: "('a => 'b) => 'a => 'b" krauss@32462: where krauss@32462: wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==> krauss@32462: wfrec_rel R F x (F g x)" krauss@32462: haftmann@35416: definition haftmann@35416: cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" where krauss@32462: "cut f r x == (%y. if (y,x):r then f y else undefined)" krauss@32462: haftmann@35416: definition haftmann@35416: adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" where krauss@32462: "adm_wf R F == ALL f g x. krauss@32462: (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x" krauss@32462: haftmann@35416: definition haftmann@35416: wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" where krauss@32462: [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y" krauss@32462: krauss@32462: subsection{*Well-Founded Recursion*} krauss@32462: krauss@32462: text{*cut*} krauss@32462: krauss@32462: lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))" krauss@32462: by (simp add: expand_fun_eq cut_def) krauss@32462: krauss@32462: lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)" krauss@32462: by (simp add: cut_def) krauss@32462: krauss@32462: text{*Inductive characterization of wfrec combinator; for details see: krauss@32462: John Harrison, "Inductive definitions: automation and application"*} krauss@32462: krauss@32462: lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y" krauss@32462: apply (simp add: adm_wf_def) krauss@32462: apply (erule_tac a=x in wf_induct) krauss@32462: apply (rule ex1I) krauss@32462: apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI) krauss@32462: apply (fast dest!: theI') krauss@32462: apply (erule wfrec_rel.cases, simp) krauss@32462: apply (erule allE, erule allE, erule allE, erule mp) krauss@32462: apply (fast intro: the_equality [symmetric]) krauss@32462: done krauss@32462: krauss@32462: lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)" krauss@32462: apply (simp add: adm_wf_def) krauss@32462: apply (intro strip) krauss@32462: apply (rule cuts_eq [THEN iffD2, THEN subst], assumption) krauss@32462: apply (rule refl) krauss@32462: done krauss@32462: krauss@32462: lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a" krauss@32462: apply (simp add: wfrec_def) krauss@32462: apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption) krauss@32462: apply (rule wfrec_rel.wfrecI) krauss@32462: apply (intro strip) krauss@32462: apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) krauss@32462: done krauss@32462: krauss@32462: krauss@26748: text{** This form avoids giant explosions in proofs. NOTE USE OF ==*} krauss@26748: lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a" krauss@26748: apply auto krauss@26748: apply (blast intro: wfrec) krauss@26748: done krauss@26748: krauss@26748: krauss@26748: lemma tfl_wf_induct: "ALL R. wf R --> krauss@26748: (ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))" krauss@26748: apply clarify krauss@26748: apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast) krauss@26748: done krauss@26748: krauss@26748: lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)" krauss@26748: apply clarify krauss@26748: apply (rule cut_apply, assumption) krauss@26748: done krauss@26748: krauss@26748: lemma tfl_wfrec: krauss@26748: "ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)" krauss@26748: apply clarify krauss@26748: apply (erule wfrec) krauss@26748: done krauss@26748: wenzelm@10773: lemma tfl_eq_True: "(x = True) --> x" wenzelm@10773: by blast wenzelm@10773: wenzelm@10773: lemma tfl_rev_eq_mp: "(x = y) --> y --> x"; wenzelm@10773: by blast wenzelm@10773: wenzelm@10773: lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)" wenzelm@10773: by blast wenzelm@6438: wenzelm@10773: lemma tfl_P_imp_P_iff_True: "P ==> P = True" wenzelm@10773: by blast wenzelm@10773: wenzelm@10773: lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)" wenzelm@10773: by blast wenzelm@10773: wenzelm@12023: lemma tfl_disj_assoc: "(a \ b) \ c == a \ (b \ c)" wenzelm@10773: by simp wenzelm@10773: wenzelm@12023: lemma tfl_disjE: "P \ Q ==> P --> R ==> Q --> R ==> R" wenzelm@10773: by blast wenzelm@10773: wenzelm@12023: lemma tfl_exE: "\x. P x ==> \x. P x --> Q ==> Q" wenzelm@10773: by blast wenzelm@10773: wenzelm@23150: use "Tools/TFL/casesplit.ML" wenzelm@23150: use "Tools/TFL/utils.ML" wenzelm@23150: use "Tools/TFL/usyntax.ML" wenzelm@23150: use "Tools/TFL/dcterm.ML" wenzelm@23150: use "Tools/TFL/thms.ML" wenzelm@23150: use "Tools/TFL/rules.ML" wenzelm@23150: use "Tools/TFL/thry.ML" wenzelm@23150: use "Tools/TFL/tfl.ML" wenzelm@23150: use "Tools/TFL/post.ML" haftmann@31723: use "Tools/recdef.ML" haftmann@31723: setup Recdef.setup wenzelm@6438: krauss@32244: text {*Wellfoundedness of @{text same_fst}*} krauss@32244: krauss@32244: definition krauss@32244: same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set" krauss@32244: where krauss@32244: "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}" krauss@32244: --{*For @{text rec_def} declarations where the first n parameters krauss@32244: stay unchanged in the recursive call. *} krauss@32244: krauss@32244: lemma same_fstI [intro!]: krauss@32244: "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R" krauss@32244: by (simp add: same_fst_def) krauss@32244: krauss@32244: lemma wf_same_fst: krauss@32244: assumes prem: "(!!x. P x ==> wf(R x))" krauss@32244: shows "wf(same_fst P R)" krauss@32244: apply (simp cong del: imp_cong add: wf_def same_fst_def) krauss@32244: apply (intro strip) krauss@32244: apply (rename_tac a b) krauss@32244: apply (case_tac "wf (R a)") krauss@32244: apply (erule_tac a = b in wf_induct, blast) krauss@32244: apply (blast intro: prem) krauss@32244: done krauss@32244: krauss@32244: text {*Rule setup*} krauss@32244: wenzelm@9855: lemmas [recdef_simp] = wenzelm@9855: inv_image_def wenzelm@9855: measure_def wenzelm@9855: lex_prod_def nipkow@11284: same_fst_def wenzelm@9855: less_Suc_eq [THEN iffD2] wenzelm@9855: wenzelm@23150: lemmas [recdef_cong] = krauss@22622: if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong wenzelm@9855: wenzelm@9855: lemmas [recdef_wf] = wenzelm@9855: wf_trancl wenzelm@9855: wf_less_than wenzelm@9855: wf_lex_prod wenzelm@9855: wf_inv_image wenzelm@9855: wf_measure wenzelm@9855: wf_pred_nat nipkow@10653: wf_same_fst wenzelm@9855: wf_empty wenzelm@9855: wenzelm@6438: end