| author | paulson | 
| Thu, 14 Jun 2007 13:18:59 +0200 | |
| changeset 23386 | 9255c1a75ba9 | 
| parent 23203 | a5026e73cfcf | 
| child 23494 | f985f9239e0d | 
| permissions | -rw-r--r-- | 
| 20324 | 1 | (* Title: HOL/FunDef.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Alexander Krauss, TU Muenchen | |
| 22816 | 4 | *) | 
| 20324 | 5 | |
| 22816 | 6 | header {* General recursive function definitions *}
 | 
| 20324 | 7 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 8 | theory FunDef | 
| 22919 
3de2d0b5b89a
explicit import of Datatype.thy due to hook bootstrap problem
 haftmann parents: 
22838diff
changeset | 9 | imports Datatype Accessible_Part | 
| 22816 | 10 | uses | 
| 11 |   ("Tools/function_package/sum_tools.ML")
 | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
22919diff
changeset | 12 |   ("Tools/function_package/fundef_lib.ML")
 | 
| 22816 | 13 |   ("Tools/function_package/fundef_common.ML")
 | 
| 14 |   ("Tools/function_package/inductive_wrap.ML")
 | |
| 15 |   ("Tools/function_package/context_tree.ML")
 | |
| 16 |   ("Tools/function_package/fundef_core.ML")
 | |
| 17 |   ("Tools/function_package/mutual.ML")
 | |
| 18 |   ("Tools/function_package/pattern_split.ML")
 | |
| 19 |   ("Tools/function_package/fundef_package.ML")
 | |
| 20 |   ("Tools/function_package/auto_term.ML")
 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 21 | begin | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 22 | |
| 22816 | 23 | text {* Definitions with default value. *}
 | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 24 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 25 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21364diff
changeset | 26 |   THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where
 | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 27 | "THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)" | 
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 28 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 29 | lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)" | 
| 22816 | 30 | by (simp add: theI' THE_default_def) | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 31 | |
| 22816 | 32 | lemma THE_default1_equality: | 
| 33 | "\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> THE_default d P = a" | |
| 34 | by (simp add: the1_equality THE_default_def) | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 35 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 36 | lemma THE_default_none: | 
| 22816 | 37 | "\<not>(\<exists>!x. P x) \<Longrightarrow> THE_default d P = d" | 
| 38 | by (simp add:THE_default_def) | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 39 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 40 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 41 | lemma fundef_ex1_existence: | 
| 22816 | 42 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 43 | assumes ex1: "\<exists>!y. G x y" | |
| 44 | shows "G x (f x)" | |
| 45 | apply (simp only: f_def) | |
| 46 | apply (rule THE_defaultI') | |
| 47 | apply (rule ex1) | |
| 48 | done | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 49 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 50 | lemma fundef_ex1_uniqueness: | 
| 22816 | 51 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 52 | assumes ex1: "\<exists>!y. G x y" | |
| 53 | assumes elm: "G x (h x)" | |
| 54 | shows "h x = f x" | |
| 55 | apply (simp only: f_def) | |
| 56 | apply (rule THE_default1_equality [symmetric]) | |
| 57 | apply (rule ex1) | |
| 58 | apply (rule elm) | |
| 59 | done | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 60 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 61 | lemma fundef_ex1_iff: | 
| 22816 | 62 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 63 | assumes ex1: "\<exists>!y. G x y" | |
| 64 | shows "(G x y) = (f x = y)" | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 65 | apply (auto simp:ex1 f_def THE_default1_equality) | 
| 22816 | 66 | apply (rule THE_defaultI') | 
| 67 | apply (rule ex1) | |
| 68 | done | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 69 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 70 | lemma fundef_default_value: | 
| 22816 | 71 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 72 | assumes graph: "\<And>x y. G x y \<Longrightarrow> D x" | |
| 73 | assumes "\<not> D x" | |
| 74 | shows "f x = d x" | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 75 | proof - | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 76 | have "\<not>(\<exists>y. G x y)" | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 77 | proof | 
| 21512 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 78 | assume "\<exists>y. G x y" | 
| 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 79 | hence "D x" using graph .. | 
| 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 80 | with `\<not> D x` show False .. | 
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 81 | qed | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 82 | hence "\<not>(\<exists>!y. G x y)" by blast | 
| 22816 | 83 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 84 | thus ?thesis | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 85 | unfolding f_def | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 86 | by (rule THE_default_none) | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 87 | qed | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 88 | |
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 89 | |
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 90 | use "Tools/function_package/sum_tools.ML" | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
22919diff
changeset | 91 | use "Tools/function_package/fundef_lib.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 92 | use "Tools/function_package/fundef_common.ML" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20324diff
changeset | 93 | use "Tools/function_package/inductive_wrap.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 94 | use "Tools/function_package/context_tree.ML" | 
| 22166 | 95 | use "Tools/function_package/fundef_core.ML" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 96 | use "Tools/function_package/mutual.ML" | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 97 | use "Tools/function_package/pattern_split.ML" | 
| 21319 
cf814e36f788
replaced "auto_term" by the simpler method "relation", which does not try
 krauss parents: 
21312diff
changeset | 98 | use "Tools/function_package/auto_term.ML" | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 99 | use "Tools/function_package/fundef_package.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 100 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 101 | setup FundefPackage.setup | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 102 | |
| 22838 | 103 | lemma let_cong [fundef_cong]: | 
| 104 | "M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g" | |
| 22816 | 105 | unfolding Let_def by blast | 
| 22622 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22325diff
changeset | 106 | |
| 22816 | 107 | lemmas [fundef_cong] = | 
| 22838 | 108 | if_cong image_cong INT_cong UN_cong | 
| 109 | bex_cong ball_cong imp_cong | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 110 | |
| 22816 | 111 | lemma split_cong [fundef_cong]: | 
| 22838 | 112 | "(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q | 
| 22816 | 113 | \<Longrightarrow> split f p = split g q" | 
| 114 | by (auto simp: split_def) | |
| 19934 | 115 | |
| 22816 | 116 | lemma comp_cong [fundef_cong]: | 
| 22838 | 117 | "f (g x) = f' (g' x') \<Longrightarrow> (f o g) x = (f' o g') x'" | 
| 22816 | 118 | unfolding o_apply . | 
| 19934 | 119 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 120 | end |