| author | immler | 
| Mon, 07 Jan 2019 18:50:41 +0100 | |
| changeset 69622 | 003475955593 | 
| parent 69605 | a96320074298 | 
| child 69913 | ca515cf61651 | 
| 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  | 
|
| 60500 | 5  | 
section \<open>TFL: recursive function definitions\<close>  | 
| 7701 | 6  | 
|
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
7  | 
theory Old_Recdef  | 
| 55017 | 8  | 
imports Main  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
9  | 
keywords  | 
| 60523 | 10  | 
"recdef" :: thy_decl and  | 
| 
46950
 
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
 
wenzelm 
parents: 
46947 
diff
changeset
 | 
11  | 
"permissive" "congs" "hints"  | 
| 15131 | 12  | 
begin  | 
| 10773 | 13  | 
|
| 60500 | 14  | 
subsection \<open>Lemmas for TFL\<close>  | 
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
15  | 
|
| 67091 | 16  | 
lemma tfl_wf_induct: "\<forall>R. wf R \<longrightarrow>  | 
17  | 
(\<forall>P. (\<forall>x. (\<forall>y. (y,x)\<in>R \<longrightarrow> P y) \<longrightarrow> P x) \<longrightarrow> (\<forall>x. P x))"  | 
|
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
18  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
19  | 
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
 | 
20  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
21  | 
|
| 
58184
 
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
 
hoelzl 
parents: 
56248 
diff
changeset
 | 
22  | 
lemma tfl_cut_def: "cut f r x \<equiv> (\<lambda>y. if (y,x) \<in> r then f y else undefined)"  | 
| 
 
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
 
hoelzl 
parents: 
56248 
diff
changeset
 | 
23  | 
unfolding cut_def .  | 
| 
 
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
 
hoelzl 
parents: 
56248 
diff
changeset
 | 
24  | 
|
| 67091 | 25  | 
lemma tfl_cut_apply: "\<forall>f R. (x,a)\<in>R \<longrightarrow> (cut f R a)(x) = f(x)"  | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
26  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
27  | 
apply (rule cut_apply, assumption)  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
28  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
29  | 
|
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
30  | 
lemma tfl_wfrec:  | 
| 67091 | 31  | 
"\<forall>M R f. (f=wfrec R M) \<longrightarrow> wf R \<longrightarrow> (\<forall>x. f x = M (cut f R x) x)"  | 
| 
26748
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
32  | 
apply clarify  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
33  | 
apply (erule wfrec)  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
34  | 
done  | 
| 
 
4d51ddd6aa5c
Merged theories about wellfoundedness into one: Wellfounded.thy
 
krauss 
parents: 
23742 
diff
changeset
 | 
35  | 
|
| 67091 | 36  | 
lemma tfl_eq_True: "(x = True) \<longrightarrow> x"  | 
| 10773 | 37  | 
by blast  | 
38  | 
||
| 67091 | 39  | 
lemma tfl_rev_eq_mp: "(x = y) \<longrightarrow> y \<longrightarrow> x"  | 
| 10773 | 40  | 
by blast  | 
41  | 
||
| 67091 | 42  | 
lemma tfl_simp_thm: "(x \<longrightarrow> y) \<longrightarrow> (x = x') \<longrightarrow> (x' \<longrightarrow> y)"  | 
| 10773 | 43  | 
by blast  | 
| 6438 | 44  | 
|
| 67091 | 45  | 
lemma tfl_P_imp_P_iff_True: "P \<Longrightarrow> P = True"  | 
| 10773 | 46  | 
by blast  | 
47  | 
||
| 67091 | 48  | 
lemma tfl_imp_trans: "(A \<longrightarrow> B) \<Longrightarrow> (B \<longrightarrow> C) \<Longrightarrow> (A \<longrightarrow> C)"  | 
| 10773 | 49  | 
by blast  | 
50  | 
||
| 67091 | 51  | 
lemma tfl_disj_assoc: "(a \<or> b) \<or> c \<equiv> a \<or> (b \<or> c)"  | 
| 10773 | 52  | 
by simp  | 
53  | 
||
| 67091 | 54  | 
lemma tfl_disjE: "P \<or> Q \<Longrightarrow> P \<longrightarrow> R \<Longrightarrow> Q \<longrightarrow> R \<Longrightarrow> R"  | 
| 10773 | 55  | 
by blast  | 
56  | 
||
| 67091 | 57  | 
lemma tfl_exE: "\<exists>x. P x \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q \<Longrightarrow> Q"  | 
| 10773 | 58  | 
by blast  | 
59  | 
||
| 69605 | 60  | 
ML_file \<open>old_recdef.ML\<close>  | 
| 58825 | 61  | 
|
| 6438 | 62  | 
|
| 60500 | 63  | 
subsection \<open>Rule setup\<close>  | 
| 32244 | 64  | 
|
| 9855 | 65  | 
lemmas [recdef_simp] =  | 
66  | 
inv_image_def  | 
|
67  | 
measure_def  | 
|
68  | 
lex_prod_def  | 
|
| 11284 | 69  | 
same_fst_def  | 
| 9855 | 70  | 
less_Suc_eq [THEN iffD2]  | 
71  | 
||
| 23150 | 72  | 
lemmas [recdef_cong] =  | 
| 
56248
 
67dc9549fa15
generalized and strengthened cong rules on compound operators, similar to 1ed737a98198
 
haftmann 
parents: 
55017 
diff
changeset
 | 
73  | 
if_cong let_cong image_cong INF_cong SUP_cong bex_cong ball_cong imp_cong  | 
| 60520 | 74  | 
map_cong filter_cong takeWhile_cong dropWhile_cong foldl_cong foldr_cong  | 
| 9855 | 75  | 
|
76  | 
lemmas [recdef_wf] =  | 
|
77  | 
wf_trancl  | 
|
78  | 
wf_less_than  | 
|
79  | 
wf_lex_prod  | 
|
80  | 
wf_inv_image  | 
|
81  | 
wf_measure  | 
|
| 
44013
 
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
 
krauss 
parents: 
39302 
diff
changeset
 | 
82  | 
wf_measures  | 
| 9855 | 83  | 
wf_pred_nat  | 
| 10653 | 84  | 
wf_same_fst  | 
| 9855 | 85  | 
wf_empty  | 
86  | 
||
| 6438 | 87  | 
end  |