author | paulson |
Fri, 12 Jul 2002 11:24:40 +0200 | |
changeset 13352 | 3cd767f8d78b |
parent 12437 | 6d4e02b6dd43 |
child 15131 | c69542757a4d |
permissions | -rw-r--r-- |
7701 | 1 |
(* Title: HOL/Recdef.thy |
2 |
ID: $Id$ |
|
10773 | 3 |
Author: Konrad Slind and Markus Wenzel, TU Muenchen |
12023 | 4 |
*) |
5123 | 5 |
|
12023 | 6 |
header {* TFL: recursive function definitions *} |
7701 | 7 |
|
10212 | 8 |
theory Recdef = Wellfounded_Relations + Datatype |
7701 | 9 |
files |
10773 | 10 |
("../TFL/utils.ML") |
11 |
("../TFL/usyntax.ML") |
|
12 |
("../TFL/dcterm.ML") |
|
13 |
("../TFL/thms.ML") |
|
14 |
("../TFL/rules.ML") |
|
15 |
("../TFL/thry.ML") |
|
16 |
("../TFL/tfl.ML") |
|
17 |
("../TFL/post.ML") |
|
12437
6d4e02b6dd43
Moved code generator setup from Recdef to Inductive.
berghofe
parents:
12023
diff
changeset
|
18 |
("Tools/recdef_package.ML"): |
10773 | 19 |
|
20 |
lemma tfl_eq_True: "(x = True) --> x" |
|
21 |
by blast |
|
22 |
||
23 |
lemma tfl_rev_eq_mp: "(x = y) --> y --> x"; |
|
24 |
by blast |
|
25 |
||
26 |
lemma tfl_simp_thm: "(x --> y) --> (x = x') --> (x' --> y)" |
|
27 |
by blast |
|
6438 | 28 |
|
10773 | 29 |
lemma tfl_P_imp_P_iff_True: "P ==> P = True" |
30 |
by blast |
|
31 |
||
32 |
lemma tfl_imp_trans: "(A --> B) ==> (B --> C) ==> (A --> C)" |
|
33 |
by blast |
|
34 |
||
12023 | 35 |
lemma tfl_disj_assoc: "(a \<or> b) \<or> c == a \<or> (b \<or> c)" |
10773 | 36 |
by simp |
37 |
||
12023 | 38 |
lemma tfl_disjE: "P \<or> Q ==> P --> R ==> Q --> R ==> R" |
10773 | 39 |
by blast |
40 |
||
12023 | 41 |
lemma tfl_exE: "\<exists>x. P x ==> \<forall>x. P x --> Q ==> Q" |
10773 | 42 |
by blast |
43 |
||
44 |
use "../TFL/utils.ML" |
|
45 |
use "../TFL/usyntax.ML" |
|
46 |
use "../TFL/dcterm.ML" |
|
47 |
use "../TFL/thms.ML" |
|
48 |
use "../TFL/rules.ML" |
|
49 |
use "../TFL/thry.ML" |
|
50 |
use "../TFL/tfl.ML" |
|
51 |
use "../TFL/post.ML" |
|
52 |
use "Tools/recdef_package.ML" |
|
6438 | 53 |
setup RecdefPackage.setup |
54 |
||
9855 | 55 |
lemmas [recdef_simp] = |
56 |
inv_image_def |
|
57 |
measure_def |
|
58 |
lex_prod_def |
|
11284 | 59 |
same_fst_def |
9855 | 60 |
less_Suc_eq [THEN iffD2] |
61 |
||
11165 | 62 |
lemmas [recdef_cong] = if_cong image_cong |
9855 | 63 |
|
64 |
lemma let_cong [recdef_cong]: |
|
65 |
"M = N ==> (!!x. x = N ==> f x = g x) ==> Let M f = Let N g" |
|
66 |
by (unfold Let_def) blast |
|
67 |
||
68 |
lemmas [recdef_wf] = |
|
69 |
wf_trancl |
|
70 |
wf_less_than |
|
71 |
wf_lex_prod |
|
72 |
wf_inv_image |
|
73 |
wf_measure |
|
74 |
wf_pred_nat |
|
10653 | 75 |
wf_same_fst |
9855 | 76 |
wf_empty |
77 |
||
6438 | 78 |
end |