| author | wenzelm | 
| Mon, 26 Mar 2012 21:03:30 +0200 | |
| changeset 47133 | 89b13238d7f2 | 
| parent 46950 | d0181abdbdac | 
| child 48891 | c0eafbd55de3 | 
| permissions | -rw-r--r-- | 
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Old_Recdef.thy  | 
| 10773 | 2  | 
Author: Konrad Slind and Markus Wenzel, TU Muenchen  | 
| 12023 | 3  | 
*)  | 
| 5123 | 4  | 
|
| 12023 | 5  | 
header {* TFL: recursive function definitions *}
 | 
| 7701 | 6  | 
|
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
7  | 
theory Old_Recdef  | 
| 
44014
 
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
 
krauss 
parents: 
44013 
diff
changeset
 | 
8  | 
imports Wfrec  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
9  | 
keywords  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
10  | 
"recdef" "defer_recdef" :: thy_decl and  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
11  | 
"recdef_tc" :: thy_goal and  | 
| 
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
12  | 
"permissive" "congs" "hints"  | 
| 16417 | 13  | 
uses  | 
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
14  | 
  ("~~/src/HOL/Tools/TFL/casesplit.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
15  | 
  ("~~/src/HOL/Tools/TFL/utils.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
16  | 
  ("~~/src/HOL/Tools/TFL/usyntax.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
17  | 
  ("~~/src/HOL/Tools/TFL/dcterm.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
18  | 
  ("~~/src/HOL/Tools/TFL/thms.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
19  | 
  ("~~/src/HOL/Tools/TFL/rules.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
20  | 
  ("~~/src/HOL/Tools/TFL/thry.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
21  | 
  ("~~/src/HOL/Tools/TFL/tfl.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
22  | 
  ("~~/src/HOL/Tools/TFL/post.ML")
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
23  | 
  ("~~/src/HOL/Tools/recdef.ML")
 | 
| 15131 | 24  | 
begin  | 
| 10773 | 25  | 
|
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
26  | 
subsection {* Lemmas for TFL *}
 | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
27  | 
|
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
28  | 
lemma tfl_wf_induct: "ALL R. wf R -->  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
29  | 
(ALL P. (ALL x. (ALL y. (y,x):R --> P y) --> P x) --> (ALL x. P x))"  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
30  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
31  | 
apply (rule_tac r = R and P = P and a = x in wf_induct, assumption, blast)  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
32  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
33  | 
|
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
34  | 
lemma tfl_cut_apply: "ALL f R. (x,a):R --> (cut f R a)(x) = f(x)"  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
35  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
36  | 
apply (rule cut_apply, assumption)  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
37  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
38  | 
|
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
39  | 
lemma tfl_wfrec:  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
40  | 
"ALL M R f. (f=wfrec R M) --> wf R --> (ALL x. f x = M (cut f R x) x)"  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
41  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
42  | 
apply (erule wfrec)  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
43  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
44  | 
|
| 10773 | 45  | 
lemma tfl_eq_True: "(x = True) --> x"  | 
46  | 
by blast  | 
|
47  | 
||
| 46947 | 48  | 
lemma tfl_rev_eq_mp: "(x = y) --> y --> x"  | 
| 10773 | 49  | 
by blast  | 
50  | 
||
51  | 
lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)"  | 
|
52  | 
by blast  | 
|
| 6438 | 53  | 
|
| 10773 | 54  | 
lemma tfl_P_imp_P_iff_True: "P ==> P = True"  | 
55  | 
by blast  | 
|
56  | 
||
57  | 
lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)"  | 
|
58  | 
by blast  | 
|
59  | 
||
| 12023 | 60  | 
lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)"  | 
| 10773 | 61  | 
by simp  | 
62  | 
||
| 12023 | 63  | 
lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R"  | 
| 10773 | 64  | 
by blast  | 
65  | 
||
| 12023 | 66  | 
lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q"  | 
| 10773 | 67  | 
by blast  | 
68  | 
||
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
69  | 
use "~~/src/HOL/Tools/TFL/casesplit.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
70  | 
use "~~/src/HOL/Tools/TFL/utils.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
71  | 
use "~~/src/HOL/Tools/TFL/usyntax.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
72  | 
use "~~/src/HOL/Tools/TFL/dcterm.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
73  | 
use "~~/src/HOL/Tools/TFL/thms.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
74  | 
use "~~/src/HOL/Tools/TFL/rules.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
75  | 
use "~~/src/HOL/Tools/TFL/thry.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
76  | 
use "~~/src/HOL/Tools/TFL/tfl.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
77  | 
use "~~/src/HOL/Tools/TFL/post.ML"  | 
| 
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
78  | 
use "~~/src/HOL/Tools/recdef.ML"  | 
| 
31723
 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 
haftmann 
parents: 
29654 
diff
changeset
 | 
79  | 
setup Recdef.setup  | 
| 6438 | 80  | 
|
| 
44014
 
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
 
krauss 
parents: 
44013 
diff
changeset
 | 
81  | 
subsection {* Rule setup *}
 | 
| 32244 | 82  | 
|
| 9855 | 83  | 
lemmas [recdef_simp] =  | 
84  | 
inv_image_def  | 
|
85  | 
measure_def  | 
|
86  | 
lex_prod_def  | 
|
| 11284 | 87  | 
same_fst_def  | 
| 9855 | 88  | 
less_Suc_eq [THEN iffD2]  | 
89  | 
||
| 23150 | 90  | 
lemmas [recdef_cong] =  | 
| 
22622
 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 
krauss 
parents: 
22399 
diff
changeset
 | 
91  | 
if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong  | 
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
92  | 
map_cong filter_cong takeWhile_cong dropWhile_cong foldl_cong foldr_cong  | 
| 9855 | 93  | 
|
94  | 
lemmas [recdef_wf] =  | 
|
95  | 
wf_trancl  | 
|
96  | 
wf_less_than  | 
|
97  | 
wf_lex_prod  | 
|
98  | 
wf_inv_image  | 
|
99  | 
wf_measure  | 
|
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
100  | 
wf_measures  | 
| 9855 | 101  | 
wf_pred_nat  | 
| 10653 | 102  | 
wf_same_fst  | 
| 9855 | 103  | 
wf_empty  | 
104  | 
||
| 6438 | 105  | 
end  |