diff -r 4453cfb745e5 -r f90e3926e627 src/HOL/Fun_Def.thy --- a/src/HOL/Fun_Def.thy Wed Aug 10 22:05:00 2016 +0200 +++ b/src/HOL/Fun_Def.thy Wed Aug 10 22:05:36 2016 +0200 @@ -5,32 +5,29 @@ section \Function Definitions and Termination Proofs\ theory Fun_Def -imports Basic_BNF_LFPs Partial_Function SAT -keywords - "function" "termination" :: thy_goal and - "fun" "fun_cases" :: thy_decl + imports Basic_BNF_LFPs Partial_Function SAT + keywords + "function" "termination" :: thy_goal and + "fun" "fun_cases" :: thy_decl begin subsection \Definitions with default value\ -definition - THE_default :: "'a \ ('a \ bool) \ 'a" where - "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" +definition THE_default :: "'a \ ('a \ bool) \ 'a" + where "THE_default d P = (if (\!x. P x) then (THE x. P x) else d)" lemma THE_defaultI': "\!x. P x \ P (THE_default d P)" by (simp add: theI' THE_default_def) -lemma THE_default1_equality: - "\\!x. P x; P a\ \ THE_default d P = a" +lemma THE_default1_equality: "\!x. P x \ P a \ THE_default d P = a" by (simp add: the1_equality THE_default_def) -lemma THE_default_none: - "\(\!x. P x) \ THE_default d P = d" - by (simp add:THE_default_def) +lemma THE_default_none: "\ (\!x. P x) \ THE_default d P = d" + by (simp add: THE_default_def) lemma fundef_ex1_existence: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "G x (f x)" apply (simp only: f_def) @@ -39,7 +36,7 @@ done lemma fundef_ex1_uniqueness: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" assumes elm: "G x (h x)" shows "h x = f x" @@ -50,7 +47,7 @@ done lemma fundef_ex1_iff: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes ex1: "\!y. G x y" shows "(G x y) = (f x = y)" apply (auto simp:ex1 f_def THE_default1_equality) @@ -59,7 +56,7 @@ done lemma fundef_default_value: - assumes f_def: "f == (\x::'a. THE_default (d x) (\y. G x y))" + assumes f_def: "f \ (\x::'a. THE_default (d x) (\y. G x y))" assumes graph: "\x y. G x y \ D x" assumes "\ D x" shows "f x = d x" @@ -67,21 +64,17 @@ have "\(\y. G x y)" proof assume "\y. G x y" - hence "D x" using graph .. + then have "D x" using graph .. with \\ D x\ show False .. qed - hence "\(\!y. G x y)" by blast - - thus ?thesis - unfolding f_def - by (rule THE_default_none) + then have "\(\!y. G x y)" by blast + then show ?thesis + unfolding f_def by (rule THE_default_none) qed -definition in_rel_def[simp]: - "in_rel R x y == (x, y) \ R" +definition in_rel_def[simp]: "in_rel R x y \ (x, y) \ R" -lemma wf_in_rel: - "wf R \ wfP (in_rel R)" +lemma wf_in_rel: "wf R \ wfP (in_rel R)" by (simp add: wfP_def) ML_file "Tools/Function/function_core.ML" @@ -112,18 +105,19 @@ subsection \Measure functions\ inductive is_measure :: "('a \ nat) \ bool" -where is_measure_trivial: "is_measure f" + where is_measure_trivial: "is_measure f" named_theorems measure_function "rules that guide the heuristic generation of measure functions" ML_file "Tools/Function/measure_functions.ML" lemma measure_size[measure_function]: "is_measure size" -by (rule is_measure_trivial) + by (rule is_measure_trivial) lemma measure_fst[measure_function]: "is_measure f \ is_measure (\p. f (fst p))" -by (rule is_measure_trivial) + by (rule is_measure_trivial) + lemma measure_snd[measure_function]: "is_measure f \ is_measure (\p. f (snd p))" -by (rule is_measure_trivial) + by (rule is_measure_trivial) ML_file "Tools/Function/lexicographic_order.ML" @@ -135,8 +129,7 @@ subsection \Congruence rules\ -lemma let_cong [fundef_cong]: - "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" +lemma let_cong [fundef_cong]: "M = N \ (\x. x = N \ f x = g x) \ Let M f = Let N g" unfolding Let_def by blast lemmas [fundef_cong] = @@ -144,13 +137,11 @@ bex_cong ball_cong imp_cong map_option_cong Option.bind_cong lemma split_cong [fundef_cong]: - "(\x y. (x, y) = q \ f x y = g x y) \ p = q - \ case_prod f p = case_prod g q" + "(\x y. (x, y) = q \ f x y = g x y) \ p = q \ case_prod f p = case_prod g q" by (auto simp: split_def) -lemma comp_cong [fundef_cong]: - "f (g x) = f' (g' x') \ (f o g) x = (f' o g') x'" - unfolding o_apply . +lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \ (f \ g) x = (f' \ g') x'" + by (simp only: o_apply) subsection \Simp rules for termination proofs\ @@ -163,31 +154,25 @@ less_imp_le_nat[termination_simp] le_imp_less_Suc[termination_simp] -lemma size_prod_simp[termination_simp]: - "size_prod f g p = f (fst p) + g (snd p) + Suc 0" -by (induct p) auto +lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0" + by (induct p) auto subsection \Decomposition\ -lemma less_by_empty: - "A = {} \ A \ B" -and union_comp_emptyL: - "\ A O C = {}; B O C = {} \ \ (A \ B) O C = {}" -and union_comp_emptyR: - "\ A O B = {}; A O C = {} \ \ A O (B \ C) = {}" -and wf_no_loop: - "R O R = {} \ wf R" -by (auto simp add: wf_comp_self[of R]) +lemma less_by_empty: "A = {} \ A \ B" + and union_comp_emptyL: "A O C = {} \ B O C = {} \ (A \ B) O C = {}" + and union_comp_emptyR: "A O B = {} \ A O C = {} \ A O (B \ C) = {}" + and wf_no_loop: "R O R = {} \ wf R" + by (auto simp add: wf_comp_self [of R]) subsection \Reduction pairs\ -definition - "reduction_pair P = (wf (fst P) \ fst P O snd P \ fst P)" +definition "reduction_pair P \ wf (fst P) \ fst P O snd P \ fst P" lemma reduction_pairI[intro]: "wf R \ R O S \ R \ reduction_pair (R, S)" -unfolding reduction_pair_def by auto + by (auto simp: reduction_pair_def) lemma reduction_pair_lemma: assumes rp: "reduction_pair P" @@ -204,13 +189,10 @@ ultimately show ?thesis by (rule wf_subset) qed -definition - "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" +definition "rp_inv_image = (\(R,S) f. (inv_image R f, inv_image S f))" -lemma rp_inv_image_rp: - "reduction_pair P \ reduction_pair (rp_inv_image P f)" - unfolding reduction_pair_def rp_inv_image_def split_def - by force +lemma rp_inv_image_rp: "reduction_pair P \ reduction_pair (rp_inv_image P f)" + unfolding reduction_pair_def rp_inv_image_def split_def by force subsection \Concrete orders for SCNP termination proofs\ @@ -230,70 +212,70 @@ and pair_leqI2: "a \ b \ s \ t \ ((a, s), (b, t)) \ pair_leq" and pair_lessI1: "a < b \ ((a, s), (b, t)) \ pair_less" and pair_lessI2: "a \ b \ s < t \ ((a, s), (b, t)) \ pair_less" - unfolding pair_leq_def pair_less_def by auto + by (auto simp: pair_leq_def pair_less_def) text \Introduction rules for max\ -lemma smax_emptyI: - "finite Y \ Y \ {} \ ({}, Y) \ max_strict" +lemma smax_emptyI: "finite Y \ Y \ {} \ ({}, Y) \ max_strict" and smax_insertI: - "\y \ Y; (x, y) \ pair_less; (X, Y) \ max_strict\ \ (insert x X, Y) \ max_strict" - and wmax_emptyI: - "finite X \ ({}, X) \ max_weak" + "y \ Y \ (x, y) \ pair_less \ (X, Y) \ max_strict \ (insert x X, Y) \ max_strict" + and wmax_emptyI: "finite X \ ({}, X) \ max_weak" and wmax_insertI: - "\y \ YS; (x, y) \ pair_leq; (XS, YS) \ max_weak\ \ (insert x XS, YS) \ max_weak" -unfolding max_strict_def max_weak_def by (auto elim!: max_ext.cases) + "y \ YS \ (x, y) \ pair_leq \ (XS, YS) \ max_weak \ (insert x XS, YS) \ max_weak" + by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases) text \Introduction rules for min\ -lemma smin_emptyI: - "X \ {} \ (X, {}) \ min_strict" +lemma smin_emptyI: "X \ {} \ (X, {}) \ min_strict" and smin_insertI: - "\x \ XS; (x, y) \ pair_less; (XS, YS) \ min_strict\ \ (XS, insert y YS) \ min_strict" - and wmin_emptyI: - "(X, {}) \ min_weak" + "x \ XS \ (x, y) \ pair_less \ (XS, YS) \ min_strict \ (XS, insert y YS) \ min_strict" + and wmin_emptyI: "(X, {}) \ min_weak" and wmin_insertI: - "\x \ XS; (x, y) \ pair_leq; (XS, YS) \ min_weak\ \ (XS, insert y YS) \ min_weak" -by (auto simp: min_strict_def min_weak_def min_ext_def) + "x \ XS \ (x, y) \ pair_leq \ (XS, YS) \ min_weak \ (XS, insert y YS) \ min_weak" + by (auto simp: min_strict_def min_weak_def min_ext_def) -text \Reduction Pairs\ +text \Reduction Pairs.\ lemma max_ext_compat: assumes "R O S \ R" - shows "max_ext R O (max_ext S \ {({},{})}) \ max_ext R" -using assms -apply auto -apply (elim max_ext.cases) -apply rule -apply auto[3] -apply (drule_tac x=xa in meta_spec) -apply simp -apply (erule bexE) -apply (drule_tac x=xb in meta_spec) -by auto + shows "max_ext R O (max_ext S \ {({}, {})}) \ max_ext R" + using assms + apply auto + apply (elim max_ext.cases) + apply rule + apply auto[3] + apply (drule_tac x=xa in meta_spec) + apply simp + apply (erule bexE) + apply (drule_tac x=xb in meta_spec) + apply auto + done lemma max_rpair_set: "reduction_pair (max_strict, max_weak)" unfolding max_strict_def max_weak_def -apply (intro reduction_pairI max_ext_wf) -apply simp -apply (rule max_ext_compat) -by (auto simp: pair_less_def pair_leq_def) + apply (intro reduction_pairI max_ext_wf) + apply simp + apply (rule max_ext_compat) + apply (auto simp: pair_less_def pair_leq_def) + done lemma min_ext_compat: assumes "R O S \ R" shows "min_ext R O (min_ext S \ {({},{})}) \ min_ext R" -using assms -apply (auto simp: min_ext_def) -apply (drule_tac x=ya in bspec, assumption) -apply (erule bexE) -apply (drule_tac x=xc in bspec) -apply assumption -by auto + using assms + apply (auto simp: min_ext_def) + apply (drule_tac x=ya in bspec, assumption) + apply (erule bexE) + apply (drule_tac x=xc in bspec) + apply assumption + apply auto + done lemma min_rpair_set: "reduction_pair (min_strict, min_weak)" unfolding min_strict_def min_weak_def -apply (intro reduction_pairI min_ext_wf) -apply simp -apply (rule min_ext_compat) -by (auto simp: pair_less_def pair_leq_def) + apply (intro reduction_pairI min_ext_wf) + apply simp + apply (rule min_ext_compat) + apply (auto simp: pair_less_def pair_leq_def) + done subsection \Tool setup\