1(*  Title:      HOL/Fun_Def.thy
2    Author:     Alexander Krauss, TU Muenchen
3*)
4
5section \<open>Function Definitions and Termination Proofs\<close>
6
7theory Fun_Def
8  imports Basic_BNF_LFPs Partial_Function SAT
9  keywords
10    "function" "termination" :: thy_goal and
11    "fun" "fun_cases" :: thy_decl
12begin
13
14subsection \<open>Definitions with default value\<close>
15
16definition THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a"
17  where "THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)"
18
19lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)"
20  by (simp add: theI' THE_default_def)
21
22lemma THE_default1_equality: "\<exists>!x. P x \<Longrightarrow> P a \<Longrightarrow> THE_default d P = a"
23  by (simp add: the1_equality THE_default_def)
24
25lemma THE_default_none: "\<not> (\<exists>!x. P x) \<Longrightarrow> THE_default d P = d"
26  by (simp add: THE_default_def)
27
28
29lemma fundef_ex1_existence:
30  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
31  assumes ex1: "\<exists>!y. G x y"
32  shows "G x (f x)"
33  apply (simp only: f_def)
34  apply (rule THE_defaultI')
35  apply (rule ex1)
36  done
37
38lemma fundef_ex1_uniqueness:
39  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
40  assumes ex1: "\<exists>!y. G x y"
41  assumes elm: "G x (h x)"
42  shows "h x = f x"
43  apply (simp only: f_def)
44  apply (rule THE_default1_equality [symmetric])
45   apply (rule ex1)
46  apply (rule elm)
47  done
48
49lemma fundef_ex1_iff:
50  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
51  assumes ex1: "\<exists>!y. G x y"
52  shows "(G x y) = (f x = y)"
53  apply (auto simp:ex1 f_def THE_default1_equality)
54  apply (rule THE_defaultI')
55  apply (rule ex1)
56  done
57
58lemma fundef_default_value:
59  assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))"
60  assumes graph: "\<And>x y. G x y \<Longrightarrow> D x"
61  assumes "\<not> D x"
62  shows "f x = d x"
63proof -
64  have "\<not>(\<exists>y. G x y)"
65  proof
66    assume "\<exists>y. G x y"
67    then have "D x" using graph ..
68    with \<open>\<not> D x\<close> show False ..
69  qed
70  then have "\<not>(\<exists>!y. G x y)" by blast
71  then show ?thesis
72    unfolding f_def by (rule THE_default_none)
73qed
74
75definition in_rel_def[simp]: "in_rel R x y \<equiv> (x, y) \<in> R"
76
77lemma wf_in_rel: "wf R \<Longrightarrow> wfP (in_rel R)"
78  by (simp add: wfP_def)
79
80ML_file "Tools/Function/function_core.ML"
81ML_file "Tools/Function/mutual.ML"
82ML_file "Tools/Function/pattern_split.ML"
83ML_file "Tools/Function/relation.ML"
84ML_file "Tools/Function/function_elims.ML"
85
86method_setup relation = \<open>
87  Args.term >> (fn t => fn ctxt => SIMPLE_METHOD' (Function_Relation.relation_infer_tac ctxt t))
88\<close> "prove termination using a user-specified wellfounded relation"
89
90ML_file "Tools/Function/function.ML"
91ML_file "Tools/Function/pat_completeness.ML"
92
93method_setup pat_completeness = \<open>
94  Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac)
95\<close> "prove completeness of (co)datatype patterns"
96
97ML_file "Tools/Function/fun.ML"
98ML_file "Tools/Function/induction_schema.ML"
99
100method_setup induction_schema = \<open>
101  Scan.succeed (Method.CONTEXT_TACTIC oo Induction_Schema.induction_schema_tac)
102\<close> "prove an induction principle"
103
104
105subsection \<open>Measure functions\<close>
106
107inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
108  where is_measure_trivial: "is_measure f"
109
110named_theorems measure_function "rules that guide the heuristic generation of measure functions"
111ML_file "Tools/Function/measure_functions.ML"
112
113lemma measure_size[measure_function]: "is_measure size"
114  by (rule is_measure_trivial)
115
116lemma measure_fst[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (fst p))"
117  by (rule is_measure_trivial)
118
119lemma measure_snd[measure_function]: "is_measure f \<Longrightarrow> is_measure (\<lambda>p. f (snd p))"
120  by (rule is_measure_trivial)
121
122ML_file "Tools/Function/lexicographic_order.ML"
123
124method_setup lexicographic_order = \<open>
125  Method.sections clasimp_modifiers >>
126  (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false))
127\<close> "termination prover for lexicographic orderings"
128
129
130subsection \<open>Congruence rules\<close>
131
132lemma let_cong [fundef_cong]: "M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g"
133  unfolding Let_def by blast
134
135lemmas [fundef_cong] =
136  if_cong image_cong
137  bex_cong ball_cong imp_cong map_option_cong Option.bind_cong
138
139lemma split_cong [fundef_cong]:
140  "(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q \<Longrightarrow> case_prod f p = case_prod g q"
141  by (auto simp: split_def)
142
143lemma comp_cong [fundef_cong]: "f (g x) = f' (g' x') \<Longrightarrow> (f \<circ> g) x = (f' \<circ> g') x'"
144  by (simp only: o_apply)
145
146
147subsection \<open>Simp rules for termination proofs\<close>
148
149declare
150  trans_less_add1[termination_simp]
151  trans_less_add2[termination_simp]
152  trans_le_add1[termination_simp]
153  trans_le_add2[termination_simp]
154  less_imp_le_nat[termination_simp]
155  le_imp_less_Suc[termination_simp]
156
157lemma size_prod_simp[termination_simp]: "size_prod f g p = f (fst p) + g (snd p) + Suc 0"
158  by (induct p) auto
159
160
161subsection \<open>Decomposition\<close>
162
163lemma less_by_empty: "A = {} \<Longrightarrow> A \<subseteq> B"
164  and union_comp_emptyL: "A O C = {} \<Longrightarrow> B O C = {} \<Longrightarrow> (A \<union> B) O C = {}"
165  and union_comp_emptyR: "A O B = {} \<Longrightarrow> A O C = {} \<Longrightarrow> A O (B \<union> C) = {}"
166  and wf_no_loop: "R O R = {} \<Longrightarrow> wf R"
167  by (auto simp add: wf_comp_self [of R])
168
169
170subsection \<open>Reduction pairs\<close>
171
172definition "reduction_pair P \<longleftrightarrow> wf (fst P) \<and> fst P O snd P \<subseteq> fst P"
173
174lemma reduction_pairI[intro]: "wf R \<Longrightarrow> R O S \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
175  by (auto simp: reduction_pair_def)
176
177lemma reduction_pair_lemma:
178  assumes rp: "reduction_pair P"
179  assumes "R \<subseteq> fst P"
180  assumes "S \<subseteq> snd P"
181  assumes "wf S"
182  shows "wf (R \<union> S)"
183proof -
184  from rp \<open>S \<subseteq> snd P\<close> have "wf (fst P)" "fst P O S \<subseteq> fst P"
185    unfolding reduction_pair_def by auto
186  with \<open>wf S\<close> have "wf (fst P \<union> S)"
187    by (auto intro: wf_union_compatible)
188  moreover from \<open>R \<subseteq> fst P\<close> have "R \<union> S \<subseteq> fst P \<union> S" by auto
189  ultimately show ?thesis by (rule wf_subset)
190qed
191
192definition "rp_inv_image = (\<lambda>(R,S) f. (inv_image R f, inv_image S f))"
193
194lemma rp_inv_image_rp: "reduction_pair P \<Longrightarrow> reduction_pair (rp_inv_image P f)"
195  unfolding reduction_pair_def rp_inv_image_def split_def by force
196
197
198subsection \<open>Concrete orders for SCNP termination proofs\<close>
199
200definition "pair_less = less_than <*lex*> less_than"
201definition "pair_leq = pair_less\<^sup>="
202definition "max_strict = max_ext pair_less"
203definition "max_weak = max_ext pair_leq \<union> {({}, {})}"
204definition "min_strict = min_ext pair_less"
205definition "min_weak = min_ext pair_leq \<union> {({}, {})}"
206
207lemma wf_pair_less[simp]: "wf pair_less"
208  by (auto simp: pair_less_def)
209
210text \<open>Introduction rules for \<open>pair_less\<close>/\<open>pair_leq\<close>\<close>
211lemma pair_leqI1: "a < b \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
212  and pair_leqI2: "a \<le> b \<Longrightarrow> s \<le> t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_leq"
213  and pair_lessI1: "a < b  \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
214  and pair_lessI2: "a \<le> b \<Longrightarrow> s < t \<Longrightarrow> ((a, s), (b, t)) \<in> pair_less"
215  by (auto simp: pair_leq_def pair_less_def)
216
217text \<open>Introduction rules for max\<close>
218lemma smax_emptyI: "finite Y \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> ({}, Y) \<in> max_strict"
219  and smax_insertI:
220    "y \<in> Y \<Longrightarrow> (x, y) \<in> pair_less \<Longrightarrow> (X, Y) \<in> max_strict \<Longrightarrow> (insert x X, Y) \<in> max_strict"
221  and wmax_emptyI: "finite X \<Longrightarrow> ({}, X) \<in> max_weak"
222  and wmax_insertI:
223    "y \<in> YS \<Longrightarrow> (x, y) \<in> pair_leq \<Longrightarrow> (XS, YS) \<in> max_weak \<Longrightarrow> (insert x XS, YS) \<in> max_weak"
224  by (auto simp: max_strict_def max_weak_def elim!: max_ext.cases)
225
226text \<open>Introduction rules for min\<close>
227lemma smin_emptyI: "X \<noteq> {} \<Longrightarrow> (X, {}) \<in> min_strict"
228  and smin_insertI:
229    "x \<in> XS \<Longrightarrow> (x, y) \<in> pair_less \<Longrightarrow> (XS, YS) \<in> min_strict \<Longrightarrow> (XS, insert y YS) \<in> min_strict"
230  and wmin_emptyI: "(X, {}) \<in> min_weak"
231  and wmin_insertI:
232    "x \<in> XS \<Longrightarrow> (x, y) \<in> pair_leq \<Longrightarrow> (XS, YS) \<in> min_weak \<Longrightarrow> (XS, insert y YS) \<in> min_weak"
233  by (auto simp: min_strict_def min_weak_def min_ext_def)
234
235text \<open>Reduction Pairs.\<close>
236
237lemma max_ext_compat:
238  assumes "R O S \<subseteq> R"
239  shows "max_ext R O (max_ext S \<union> {({}, {})}) \<subseteq> max_ext R"
240  using assms
241  apply auto
242  apply (elim max_ext.cases)
243  apply rule
244     apply auto[3]
245  apply (drule_tac x=xa in meta_spec)
246  apply simp
247  apply (erule bexE)
248  apply (drule_tac x=xb in meta_spec)
249  apply auto
250  done
251
252lemma max_rpair_set: "reduction_pair (max_strict, max_weak)"
253  unfolding max_strict_def max_weak_def
254  apply (intro reduction_pairI max_ext_wf)
255   apply simp
256  apply (rule max_ext_compat)
257  apply (auto simp: pair_less_def pair_leq_def)
258  done
259
260lemma min_ext_compat:
261  assumes "R O S \<subseteq> R"
262  shows "min_ext R O  (min_ext S \<union> {({},{})}) \<subseteq> min_ext R"
263  using assms
264  apply (auto simp: min_ext_def)
265  apply (drule_tac x=ya in bspec, assumption)
266  apply (erule bexE)
267  apply (drule_tac x=xc in bspec)
268   apply assumption
269  apply auto
270  done
271
272lemma min_rpair_set: "reduction_pair (min_strict, min_weak)"
273  unfolding min_strict_def min_weak_def
274  apply (intro reduction_pairI min_ext_wf)
275   apply simp
276  apply (rule min_ext_compat)
277  apply (auto simp: pair_less_def pair_leq_def)
278  done
279
280
281subsection \<open>Yet another induction principle on the natural numbers\<close>
282
283lemma nat_descend_induct [case_names base descend]:
284  fixes P :: "nat \<Rightarrow> bool"
285  assumes H1: "\<And>k. k > n \<Longrightarrow> P k"
286  assumes H2: "\<And>k. k \<le> n \<Longrightarrow> (\<And>i. i > k \<Longrightarrow> P i) \<Longrightarrow> P k"
287  shows "P m"
288  using assms by induction_schema (force intro!: wf_measure [of "\<lambda>k. Suc n - k"])+
289
290
291subsection \<open>Tool setup\<close>
292
293ML_file "Tools/Function/termination.ML"
294ML_file "Tools/Function/scnp_solve.ML"
295ML_file "Tools/Function/scnp_reconstruct.ML"
296ML_file "Tools/Function/fun_cases.ML"
297
298ML_val \<comment> \<open>setup inactive\<close>
299\<open>
300  Context.theory_map (Function_Common.set_termination_prover
301    (K (ScnpReconstruct.decomp_scnp_tac [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])))
302\<close>
303
304end
305