blanchet@55085: (* Title: HOL/Fun_Def.thy wenzelm@20324: Author: Alexander Krauss, TU Muenchen wenzelm@22816: *) wenzelm@20324: krauss@29125: header {* Function Definitions and Termination Proofs *} wenzelm@20324: blanchet@55085: theory Fun_Def blanchet@58377: imports Basic_BNF_Least_Fixpoints Partial_Function SAT Manuel@53603: keywords "function" "termination" :: thy_goal and "fun" "fun_cases" :: thy_decl krauss@19564: begin krauss@19564: blanchet@55085: subsection {* Definitions with default value *} krauss@20536: krauss@20536: definition wenzelm@21404: THE_default :: "'a \ ('a \ bool) \ 'a" where krauss@20536: "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" krauss@20536: krauss@20536: lemma THE_defaultI': "\!x. P x \ P (THE_default d P)" wenzelm@22816: by (simp add: theI' THE_default_def) krauss@20536: wenzelm@22816: lemma THE_default1_equality: wenzelm@22816: "\\!x. P x; P a\ \ THE_default d P = a" wenzelm@22816: by (simp add: the1_equality THE_default_def) krauss@20536: krauss@20536: lemma THE_default_none: wenzelm@22816: "\(\!x. P x) \ THE_default d P = d" wenzelm@22816: by (simp add:THE_default_def) krauss@20536: krauss@20536: krauss@19564: lemma fundef_ex1_existence: wenzelm@22816: assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" wenzelm@22816: assumes ex1: "\!y. G x y" wenzelm@22816: shows "G x (f x)" wenzelm@22816: apply (simp only: f_def) wenzelm@22816: apply (rule THE_defaultI') wenzelm@22816: apply (rule ex1) wenzelm@22816: done krauss@21051: krauss@19564: lemma fundef_ex1_uniqueness: wenzelm@22816: assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" wenzelm@22816: assumes ex1: "\!y. G x y" wenzelm@22816: assumes elm: "G x (h x)" wenzelm@22816: shows "h x = f x" wenzelm@22816: apply (simp only: f_def) wenzelm@22816: apply (rule THE_default1_equality [symmetric]) wenzelm@22816: apply (rule ex1) wenzelm@22816: apply (rule elm) wenzelm@22816: done krauss@19564: krauss@19564: lemma fundef_ex1_iff: wenzelm@22816: assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" wenzelm@22816: assumes ex1: "\!y. G x y" wenzelm@22816: shows "(G x y) = (f x = y)" krauss@20536: apply (auto simp:ex1 f_def THE_default1_equality) wenzelm@22816: apply (rule THE_defaultI') wenzelm@22816: apply (rule ex1) wenzelm@22816: done krauss@19564: krauss@20654: lemma fundef_default_value: wenzelm@22816: assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" wenzelm@22816: assumes graph: "\x y. G x y \ D x" wenzelm@22816: assumes "\ D x" wenzelm@22816: shows "f x = d x" krauss@20654: proof - krauss@21051: have "\(\y. G x y)" krauss@20654: proof krauss@21512: assume "\y. G x y" krauss@21512: hence "D x" using graph .. krauss@21512: with `\ D x` show False .. krauss@20654: qed krauss@21051: hence "\(\!y. G x y)" by blast wenzelm@22816: krauss@20654: thus ?thesis krauss@20654: unfolding f_def krauss@20654: by (rule THE_default_none) krauss@20654: qed krauss@20654: berghofe@23739: definition in_rel_def[simp]: berghofe@23739: "in_rel R x y == (x, y) \ R" berghofe@23739: berghofe@23739: lemma wf_in_rel: berghofe@23739: "wf R \ wfP (in_rel R)" berghofe@23739: by (simp add: wfP_def) berghofe@23739: wenzelm@48891: ML_file "Tools/Function/function_core.ML" wenzelm@48891: ML_file "Tools/Function/mutual.ML" wenzelm@48891: ML_file "Tools/Function/pattern_split.ML" wenzelm@48891: ML_file "Tools/Function/relation.ML" Manuel@53603: ML_file "Tools/Function/function_elims.ML" wenzelm@47701: wenzelm@47701: method_setup relation = {* wenzelm@47701: Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t)) wenzelm@47701: *} "prove termination using a user-specified wellfounded relation" wenzelm@47701: wenzelm@48891: ML_file "Tools/Function/function.ML" wenzelm@48891: ML_file "Tools/Function/pat_completeness.ML" wenzelm@47432: wenzelm@47432: method_setup pat_completeness = {* wenzelm@47432: Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac) blanchet@58305: *} "prove completeness of (co)datatype patterns" wenzelm@47432: wenzelm@48891: ML_file "Tools/Function/fun.ML" wenzelm@48891: ML_file "Tools/Function/induction_schema.ML" krauss@19564: wenzelm@47432: method_setup induction_schema = {* wenzelm@58002: Scan.succeed (NO_CASES oo Induction_Schema.induction_schema_tac) wenzelm@47432: *} "prove an induction principle" wenzelm@47432: wenzelm@47701: setup {* krauss@33099: Function.setup krauss@33098: #> Function_Fun.setup krauss@25567: *} krauss@19770: blanchet@56643: blanchet@56643: subsection {* Measure functions *} krauss@29125: krauss@29125: inductive is_measure :: "('a \ nat) \ bool" krauss@29125: where is_measure_trivial: "is_measure f" krauss@29125: wenzelm@57959: named_theorems measure_function "rules that guide the heuristic generation of measure functions" wenzelm@48891: ML_file "Tools/Function/measure_functions.ML" krauss@29125: krauss@29125: lemma measure_size[measure_function]: "is_measure size" krauss@29125: by (rule is_measure_trivial) krauss@29125: krauss@29125: lemma measure_fst[measure_function]: "is_measure f \ is_measure (\p. f (fst p))" krauss@29125: by (rule is_measure_trivial) krauss@29125: lemma measure_snd[measure_function]: "is_measure f \ is_measure (\p. f (snd p))" krauss@29125: by (rule is_measure_trivial) krauss@29125: wenzelm@48891: ML_file "Tools/Function/lexicographic_order.ML" wenzelm@47432: wenzelm@47432: method_setup lexicographic_order = {* wenzelm@47432: Method.sections clasimp_modifiers >> wenzelm@47432: (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false)) wenzelm@47432: *} "termination prover for lexicographic orderings" wenzelm@47432: wenzelm@47701: setup Lexicographic_Order.setup krauss@29125: krauss@29125: blanchet@56643: subsection {* Congruence rules *} krauss@29125: haftmann@22838: lemma let_cong [fundef_cong]: haftmann@22838: "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" wenzelm@22816: unfolding Let_def by blast krauss@22622: wenzelm@22816: lemmas [fundef_cong] = haftmann@56248: if_cong image_cong INF_cong SUP_cong blanchet@55466: bex_cong ball_cong imp_cong map_option_cong Option.bind_cong krauss@19564: wenzelm@22816: lemma split_cong [fundef_cong]: haftmann@22838: "(\x y. (x, y) = q \ f x y = g x y) \ p = q wenzelm@22816: \ split f p = split g q" wenzelm@22816: by (auto simp: split_def) krauss@19934: wenzelm@22816: lemma comp_cong [fundef_cong]: haftmann@22838: "f (g x) = f' (g' x') \ (f o g) x = (f' o g') x'" wenzelm@22816: unfolding o_apply . krauss@19934: blanchet@56643: krauss@29125: subsection {* Simp rules for termination proofs *} krauss@26875: blanchet@56643: declare blanchet@56643: trans_less_add1[termination_simp] blanchet@56643: trans_less_add2[termination_simp] blanchet@56643: trans_le_add1[termination_simp] blanchet@56643: trans_le_add2[termination_simp] blanchet@56643: less_imp_le_nat[termination_simp] blanchet@56643: le_imp_less_Suc[termination_simp] krauss@26875: blanchet@56846: lemma size_prod_simp[termination_simp]: blanchet@56846: "size_prod f g p = f (fst p) + g (snd p) + Suc 0" krauss@26875: by (induct p) auto krauss@26875: blanchet@56643: krauss@29125: subsection {* Decomposition *} krauss@29125: wenzelm@47701: lemma less_by_empty: krauss@29125: "A = {} \ A \ B" krauss@29125: and union_comp_emptyL: krauss@29125: "\ A O C = {}; B O C = {} \ \ (A \ B) O C = {}" krauss@29125: and union_comp_emptyR: krauss@29125: "\ A O B = {}; A O C = {} \ \ A O (B \ C) = {}" wenzelm@47701: and wf_no_loop: krauss@29125: "R O R = {} \ wf R" krauss@29125: by (auto simp add: wf_comp_self[of R]) krauss@29125: krauss@29125: blanchet@56643: subsection {* Reduction pairs *} krauss@29125: krauss@29125: definition krauss@32235: "reduction_pair P = (wf (fst P) \ fst P O snd P \ fst P)" krauss@29125: krauss@32235: lemma reduction_pairI[intro]: "wf R \ R O S \ R \ reduction_pair (R, S)" krauss@29125: unfolding reduction_pair_def by auto krauss@29125: krauss@29125: lemma reduction_pair_lemma: krauss@29125: assumes rp: "reduction_pair P" krauss@29125: assumes "R \ fst P" krauss@29125: assumes "S \ snd P" krauss@29125: assumes "wf S" krauss@29125: shows "wf (R \ S)" krauss@29125: proof - krauss@32235: from rp `S \ snd P` have "wf (fst P)" "fst P O S \ fst P" krauss@29125: unfolding reduction_pair_def by auto wenzelm@47701: with `wf S` have "wf (fst P \ S)" krauss@29125: by (auto intro: wf_union_compatible) krauss@29125: moreover from `R \ fst P` have "R \ S \ fst P \ S" by auto wenzelm@47701: ultimately show ?thesis by (rule wf_subset) krauss@29125: qed krauss@29125: krauss@29125: definition krauss@29125: "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" krauss@29125: krauss@29125: lemma rp_inv_image_rp: krauss@29125: "reduction_pair P \ reduction_pair (rp_inv_image P f)" krauss@29125: unfolding reduction_pair_def rp_inv_image_def split_def krauss@29125: by force krauss@29125: krauss@29125: krauss@29125: subsection {* Concrete orders for SCNP termination proofs *} krauss@29125: krauss@29125: definition "pair_less = less_than <*lex*> less_than" haftmann@37767: definition "pair_leq = pair_less^=" krauss@29125: definition "max_strict = max_ext pair_less" haftmann@37767: definition "max_weak = max_ext pair_leq \ {({}, {})}" haftmann@37767: definition "min_strict = min_ext pair_less" haftmann@37767: definition "min_weak = min_ext pair_leq \ {({}, {})}" krauss@29125: krauss@29125: lemma wf_pair_less[simp]: "wf pair_less" krauss@29125: by (auto simp: pair_less_def) krauss@29125: wenzelm@29127: text {* Introduction rules for @{text pair_less}/@{text pair_leq} *} krauss@29125: lemma pair_leqI1: "a < b \ ((a, s), (b, t)) \ pair_leq" krauss@29125: and pair_leqI2: "a \ b \ s \ t \ ((a, s), (b, t)) \ pair_leq" krauss@29125: and pair_lessI1: "a < b \ ((a, s), (b, t)) \ pair_less" krauss@29125: and pair_lessI2: "a \ b \ s < t \ ((a, s), (b, t)) \ pair_less" krauss@29125: unfolding pair_leq_def pair_less_def by auto krauss@29125: krauss@29125: text {* Introduction rules for max *} wenzelm@47701: lemma smax_emptyI: wenzelm@47701: "finite Y \ Y \ {} \ ({}, Y) \ max_strict" wenzelm@47701: and smax_insertI: krauss@29125: "\y \ Y; (x, y) \ pair_less; (X, Y) \ max_strict\ \ (insert x X, Y) \ max_strict" wenzelm@47701: and wmax_emptyI: wenzelm@47701: "finite X \ ({}, X) \ max_weak" krauss@29125: and wmax_insertI: wenzelm@47701: "\y \ YS; (x, y) \ pair_leq; (XS, YS) \ max_weak\ \ (insert x XS, YS) \ max_weak" krauss@29125: unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases) krauss@29125: krauss@29125: text {* Introduction rules for min *} wenzelm@47701: lemma smin_emptyI: wenzelm@47701: "X \ {} \ (X, {}) \ min_strict" wenzelm@47701: and smin_insertI: krauss@29125: "\x \ XS; (x, y) \ pair_less; (XS, YS) \ min_strict\ \ (XS, insert y YS) \ min_strict" wenzelm@47701: and wmin_emptyI: wenzelm@47701: "(X, {}) \ min_weak" wenzelm@47701: and wmin_insertI: wenzelm@47701: "\x \ XS; (x, y) \ pair_leq; (XS, YS) \ min_weak\ \ (XS, insert y YS) \ min_weak" krauss@29125: by (auto simp: min_strict_def min_weak_def min_ext_def) krauss@29125: krauss@29125: text {* Reduction Pairs *} krauss@29125: wenzelm@47701: lemma max_ext_compat: krauss@32235: assumes "R O S \ R" krauss@32235: shows "max_ext R O (max_ext S \ {({},{})}) \ max_ext R" wenzelm@47701: using assms krauss@29125: apply auto krauss@29125: apply (elim max_ext.cases) krauss@29125: apply rule krauss@29125: apply auto[3] krauss@29125: apply (drule_tac x=xa in meta_spec) krauss@29125: apply simp krauss@29125: apply (erule bexE) krauss@29125: apply (drule_tac x=xb in meta_spec) krauss@29125: by auto krauss@29125: krauss@29125: lemma max_rpair_set: "reduction_pair (max_strict, max_weak)" wenzelm@47701: unfolding max_strict_def max_weak_def krauss@29125: apply (intro reduction_pairI max_ext_wf) krauss@29125: apply simp krauss@29125: apply (rule max_ext_compat) krauss@29125: by (auto simp: pair_less_def pair_leq_def) krauss@29125: wenzelm@47701: lemma min_ext_compat: krauss@32235: assumes "R O S \ R" krauss@32235: shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" wenzelm@47701: using assms krauss@29125: apply (auto simp: min_ext_def) krauss@29125: apply (drule_tac x=ya in bspec, assumption) krauss@29125: apply (erule bexE) krauss@29125: apply (drule_tac x=xc in bspec) krauss@29125: apply assumption krauss@29125: by auto krauss@29125: krauss@29125: lemma min_rpair_set: "reduction_pair (min_strict, min_weak)" wenzelm@47701: unfolding min_strict_def min_weak_def krauss@29125: apply (intro reduction_pairI min_ext_wf) krauss@29125: apply simp krauss@29125: apply (rule min_ext_compat) krauss@29125: by (auto simp: pair_less_def pair_leq_def) krauss@29125: krauss@29125: krauss@29125: subsection {* Tool setup *} krauss@29125: wenzelm@48891: ML_file "Tools/Function/termination.ML" wenzelm@48891: ML_file "Tools/Function/scnp_solve.ML" wenzelm@48891: ML_file "Tools/Function/scnp_reconstruct.ML" Manuel@53603: ML_file "Tools/Function/fun_cases.ML" krauss@29125: blanchet@54407: setup ScnpReconstruct.setup wenzelm@30480: wenzelm@30480: ML_val -- "setup inactive" wenzelm@30480: {* krauss@36521: Context.theory_map (Function_Common.set_termination_prover krauss@36521: (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])) krauss@29125: *} krauss@26875: krauss@19564: end