author | wenzelm |
Tue, 03 Mar 2009 14:07:43 +0100 | |
changeset 30211 | 556d1810cdad |
parent 29654 | 24e73987bfe2 |
child 31723 | f5cafe803b55 |
permissions | -rw-r--r-- |
7701 | 1 |
(* Title: HOL/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 |
|
15131 | 7 |
theory Recdef |
29654
24e73987bfe2
Plain, Main form meeting points in import hierarchy
haftmann
parents:
26748
diff
changeset
|
8 |
imports FunDef Plain |
16417 | 9 |
uses |
23150 | 10 |
("Tools/TFL/casesplit.ML") |
11 |
("Tools/TFL/utils.ML") |
|
12 |
("Tools/TFL/usyntax.ML") |
|
13 |
("Tools/TFL/dcterm.ML") |
|
14 |
("Tools/TFL/thms.ML") |
|
15 |
("Tools/TFL/rules.ML") |
|
16 |
("Tools/TFL/thry.ML") |
|
17 |
("Tools/TFL/tfl.ML") |
|
18 |
("Tools/TFL/post.ML") |
|
15131 | 19 |
("Tools/recdef_package.ML") |
20 |
begin |
|
10773 | 21 |
|
26748
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
22 |
text{** This form avoids giant explosions in proofs. NOTE USE OF ==*} |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
23 |
lemma def_wfrec: "[| f==wfrec r H; wf(r) |] ==> f(a) = H (cut f r a) a" |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
24 |
apply auto |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
25 |
apply (blast intro: wfrec) |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
26 |
done |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
27 |
|
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
28 |
|
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
29 |
lemma tfl_wf_induct: "ALL R. wf R --> |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
30 |
(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
|
31 |
apply clarify |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
32 |
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
|
33 |
done |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
34 |
|
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
35 |
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
|
36 |
apply clarify |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
37 |
apply (rule cut_apply, assumption) |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
38 |
done |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
39 |
|
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
40 |
lemma tfl_wfrec: |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
41 |
"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
|
42 |
apply clarify |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
43 |
apply (erule wfrec) |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
44 |
done |
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
krauss
parents:
23742
diff
changeset
|
45 |
|
10773 | 46 |
lemma tfl_eq_True: "(x = True) --> x" |
47 |
by blast |
|
48 |
||
49 |
lemma tfl_rev_eq_mp: "(x = y) --> y --> x"; |
|
50 |
by blast |
|
51 |
||
52 |
lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)" |
|
53 |
by blast |
|
6438 | 54 |
|
10773 | 55 |
lemma tfl_P_imp_P_iff_True: "P ==> P = True" |
56 |
by blast |
|
57 |
||
58 |
lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)" |
|
59 |
by blast |
|
60 |
||
12023 | 61 |
lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)" |
10773 | 62 |
by simp |
63 |
||
12023 | 64 |
lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R" |
10773 | 65 |
by blast |
66 |
||
12023 | 67 |
lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q" |
10773 | 68 |
by blast |
69 |
||
23150 | 70 |
use "Tools/TFL/casesplit.ML" |
71 |
use "Tools/TFL/utils.ML" |
|
72 |
use "Tools/TFL/usyntax.ML" |
|
73 |
use "Tools/TFL/dcterm.ML" |
|
74 |
use "Tools/TFL/thms.ML" |
|
75 |
use "Tools/TFL/rules.ML" |
|
76 |
use "Tools/TFL/thry.ML" |
|
77 |
use "Tools/TFL/tfl.ML" |
|
78 |
use "Tools/TFL/post.ML" |
|
10773 | 79 |
use "Tools/recdef_package.ML" |
6438 | 80 |
setup RecdefPackage.setup |
81 |
||
9855 | 82 |
lemmas [recdef_simp] = |
83 |
inv_image_def |
|
84 |
measure_def |
|
85 |
lex_prod_def |
|
11284 | 86 |
same_fst_def |
9855 | 87 |
less_Suc_eq [THEN iffD2] |
88 |
||
23150 | 89 |
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
|
90 |
if_cong let_cong image_cong INT_cong UN_cong bex_cong ball_cong imp_cong |
9855 | 91 |
|
92 |
lemmas [recdef_wf] = |
|
93 |
wf_trancl |
|
94 |
wf_less_than |
|
95 |
wf_lex_prod |
|
96 |
wf_inv_image |
|
97 |
wf_measure |
|
98 |
wf_pred_nat |
|
10653 | 99 |
wf_same_fst |
9855 | 100 |
wf_empty |
101 |
||
6438 | 102 |
end |