| author | wenzelm | 
| Wed, 07 Nov 2007 22:20:11 +0100 | |
| changeset 25332 | 73491e84ead1 | 
| parent 24699 | c6674504103f | 
| child 25556 | 8d3b7c27049b | 
| 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 | 
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24194diff
changeset | 9 | imports Accessible_Part | 
| 22816 | 10 | uses | 
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
22919diff
changeset | 11 |   ("Tools/function_package/fundef_lib.ML")
 | 
| 22816 | 12 |   ("Tools/function_package/fundef_common.ML")
 | 
| 13 |   ("Tools/function_package/inductive_wrap.ML")
 | |
| 14 |   ("Tools/function_package/context_tree.ML")
 | |
| 15 |   ("Tools/function_package/fundef_core.ML")
 | |
| 16 |   ("Tools/function_package/mutual.ML")
 | |
| 17 |   ("Tools/function_package/pattern_split.ML")
 | |
| 18 |   ("Tools/function_package/fundef_package.ML")
 | |
| 19 |   ("Tools/function_package/auto_term.ML")
 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 20 | begin | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 21 | |
| 22816 | 22 | text {* Definitions with default value. *}
 | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 23 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 24 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21364diff
changeset | 25 |   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 | 26 | "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 | 27 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 28 | lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)" | 
| 22816 | 29 | by (simp add: theI' THE_default_def) | 
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 30 | |
| 22816 | 31 | lemma THE_default1_equality: | 
| 32 | "\<lbrakk>\<exists>!x. P x; P a\<rbrakk> \<Longrightarrow> THE_default d P = a" | |
| 33 | by (simp add: the1_equality THE_default_def) | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 34 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 35 | lemma THE_default_none: | 
| 22816 | 36 | "\<not>(\<exists>!x. P x) \<Longrightarrow> THE_default d P = d" | 
| 37 | by (simp add:THE_default_def) | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 38 | |
| 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 39 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 40 | lemma fundef_ex1_existence: | 
| 22816 | 41 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 42 | assumes ex1: "\<exists>!y. G x y" | |
| 43 | shows "G x (f x)" | |
| 44 | apply (simp only: f_def) | |
| 45 | apply (rule THE_defaultI') | |
| 46 | apply (rule ex1) | |
| 47 | done | |
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 48 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 49 | lemma fundef_ex1_uniqueness: | 
| 22816 | 50 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 51 | assumes ex1: "\<exists>!y. G x y" | |
| 52 | assumes elm: "G x (h x)" | |
| 53 | shows "h x = f x" | |
| 54 | apply (simp only: f_def) | |
| 55 | apply (rule THE_default1_equality [symmetric]) | |
| 56 | apply (rule ex1) | |
| 57 | apply (rule elm) | |
| 58 | done | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 59 | |
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 60 | lemma fundef_ex1_iff: | 
| 22816 | 61 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 62 | assumes ex1: "\<exists>!y. G x y" | |
| 63 | shows "(G x y) = (f x = y)" | |
| 20536 
f088edff8af8
Function package: Outside their domain functions now return "arbitrary".
 krauss parents: 
20523diff
changeset | 64 | apply (auto simp:ex1 f_def THE_default1_equality) | 
| 22816 | 65 | apply (rule THE_defaultI') | 
| 66 | apply (rule ex1) | |
| 67 | done | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 68 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 69 | lemma fundef_default_value: | 
| 22816 | 70 | assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (\<lambda>y. G x y))" | 
| 71 | assumes graph: "\<And>x y. G x y \<Longrightarrow> D x" | |
| 72 | assumes "\<not> D x" | |
| 73 | shows "f x = d x" | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 74 | proof - | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 75 | 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 | 76 | proof | 
| 21512 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 77 | assume "\<exists>y. G x y" | 
| 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 78 | hence "D x" using graph .. | 
| 
3786eb1b69d6
Lemma "fundef_default_value" uses predicate instead of set.
 krauss parents: 
21404diff
changeset | 79 | 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 | 80 | qed | 
| 21051 
c49467a9c1e1
Switched function package to use the new package for inductive predicates.
 krauss parents: 
20654diff
changeset | 81 | hence "\<not>(\<exists>!y. G x y)" by blast | 
| 22816 | 82 | |
| 20654 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 83 | thus ?thesis | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 84 | unfolding f_def | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 85 | by (rule THE_default_none) | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 86 | qed | 
| 
d80502f0d701
1. Function package accepts a parameter (default "some_term"), which specifies the functions
 krauss parents: 
20536diff
changeset | 87 | |
| 23739 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 88 | definition in_rel_def[simp]: | 
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 89 | "in_rel R x y == (x, y) \<in> R" | 
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 90 | |
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 91 | lemma wf_in_rel: | 
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 92 | "wf R \<Longrightarrow> wfP (in_rel R)" | 
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 93 | by (simp add: wfP_def) | 
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 94 | |
| 
c5ead5df7f35
Inserted definition of in_rel again (since member2 was removed).
 berghofe parents: 
23494diff
changeset | 95 | |
| 23203 
a5026e73cfcf
"function (sequential)" and "fun" now handle incomplete patterns silently by adding "undefined" cases.
 krauss parents: 
22919diff
changeset | 96 | 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 | 97 | 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 | 98 | 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 | 99 | use "Tools/function_package/context_tree.ML" | 
| 22166 | 100 | use "Tools/function_package/fundef_core.ML" | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 101 | use "Tools/function_package/mutual.ML" | 
| 20270 
3abe7dae681e
Function package can now do automatic splits of overlapping datatype patterns
 krauss parents: 
19934diff
changeset | 102 | 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 | 103 | 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 | 104 | use "Tools/function_package/fundef_package.ML" | 
| 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 105 | |
| 24699 
c6674504103f
datatype interpretators for size and datatype_realizer
 haftmann parents: 
24194diff
changeset | 106 | setup {* FundefPackage.setup *}
 | 
| 19770 
be5c23ebe1eb
HOL/Tools/function_package: Added support for mutual recursive definitions.
 krauss parents: 
19564diff
changeset | 107 | |
| 22838 | 108 | lemma let_cong [fundef_cong]: | 
| 109 | "M = N \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x = g x) \<Longrightarrow> Let M f = Let N g" | |
| 22816 | 110 | 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 | 111 | |
| 22816 | 112 | lemmas [fundef_cong] = | 
| 22838 | 113 | if_cong image_cong INT_cong UN_cong | 
| 114 | bex_cong ball_cong imp_cong | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 115 | |
| 22816 | 116 | lemma split_cong [fundef_cong]: | 
| 22838 | 117 | "(\<And>x y. (x, y) = q \<Longrightarrow> f x y = g x y) \<Longrightarrow> p = q | 
| 22816 | 118 | \<Longrightarrow> split f p = split g q" | 
| 119 | by (auto simp: split_def) | |
| 19934 | 120 | |
| 22816 | 121 | lemma comp_cong [fundef_cong]: | 
| 22838 | 122 | "f (g x) = f' (g' x') \<Longrightarrow> (f o g) x = (f' o g') x'" | 
| 22816 | 123 | unfolding o_apply . | 
| 19934 | 124 | |
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: diff
changeset | 125 | end |