author | wenzelm |
Wed, 11 Dec 2024 10:40:57 +0100 | |
changeset 81577 | a712bf5ccab0 |
parent 75669 | 43f5dfb7fa35 |
permissions | -rw-r--r-- |
55210 | 1 |
(* Title: HOL/Wfrec.thy |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
2 |
Author: Tobias Nipkow |
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson |
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
4 |
Author: Konrad Slind |
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
5 |
*) |
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
6 |
|
60758 | 7 |
section \<open>Well-Founded Recursion Combinator\<close> |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
8 |
|
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
9 |
theory Wfrec |
63572 | 10 |
imports Wellfounded |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
11 |
begin |
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
12 |
|
63572 | 13 |
inductive wfrec_rel :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" for R F |
14 |
where wfrecI: "(\<And>z. (z, x) \<in> R \<Longrightarrow> wfrec_rel R F z (g z)) \<Longrightarrow> wfrec_rel R F x (F g x)" |
|
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
15 |
|
63572 | 16 |
definition cut :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b" |
17 |
where "cut f R x = (\<lambda>y. if (y, x) \<in> R then f y else undefined)" |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
18 |
|
63572 | 19 |
definition adm_wf :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)) \<Rightarrow> bool" |
20 |
where "adm_wf R F \<longleftrightarrow> (\<forall>f g x. (\<forall>z. (z, x) \<in> R \<longrightarrow> f z = g z) \<longrightarrow> F f x = F g x)" |
|
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
21 |
|
63572 | 22 |
definition wfrec :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)) \<Rightarrow> ('a \<Rightarrow> 'b)" |
23 |
where "wfrec R F = (\<lambda>x. THE y. wfrec_rel R (\<lambda>f x. F (cut f R x) x) x y)" |
|
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
24 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
25 |
lemma cuts_eq: "(cut f R x = cut g R x) \<longleftrightarrow> (\<forall>y. (y, x) \<in> R \<longrightarrow> f y = g y)" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
26 |
by (simp add: fun_eq_iff cut_def) |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
27 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
28 |
lemma cut_apply: "(x, a) \<in> R \<Longrightarrow> cut f R a x = f x" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
29 |
by (simp add: cut_def) |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
30 |
|
63572 | 31 |
text \<open> |
32 |
Inductive characterization of \<open>wfrec\<close> combinator; for details see: |
|
33 |
John Harrison, "Inductive definitions: automation and application". |
|
34 |
\<close> |
|
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
35 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
36 |
lemma theI_unique: "\<exists>!x. P x \<Longrightarrow> P x \<longleftrightarrow> x = The P" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
37 |
by (auto intro: the_equality[symmetric] theI) |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
38 |
|
63572 | 39 |
lemma wfrec_unique: |
40 |
assumes "adm_wf R F" "wf R" |
|
41 |
shows "\<exists>!y. wfrec_rel R F x y" |
|
60758 | 42 |
using \<open>wf R\<close> |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
43 |
proof induct |
63040 | 44 |
define f where "f y = (THE z. wfrec_rel R F y z)" for y |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
45 |
case (less x) |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
46 |
then have "\<And>y z. (y, x) \<in> R \<Longrightarrow> wfrec_rel R F y z \<longleftrightarrow> z = f y" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
47 |
unfolding f_def by (rule theI_unique) |
60758 | 48 |
with \<open>adm_wf R F\<close> show ?case |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
49 |
by (subst wfrec_rel.simps) (auto simp: adm_wf_def) |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
50 |
qed |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
51 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
52 |
lemma adm_lemma: "adm_wf R (\<lambda>f x. F (cut f R x) x)" |
63572 | 53 |
by (auto simp: adm_wf_def intro!: arg_cong[where f="\<lambda>x. F x y" for y] cuts_eq[THEN iffD2]) |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
54 |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
55 |
lemma wfrec: "wf R \<Longrightarrow> wfrec R F a = F (cut (wfrec R F) R a) a" |
63572 | 56 |
apply (simp add: wfrec_def) |
57 |
apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality]) |
|
58 |
apply assumption |
|
59 |
apply (rule wfrec_rel.wfrecI) |
|
60 |
apply (erule adm_lemma [THEN wfrec_unique, THEN theI']) |
|
61 |
done |
|
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
62 |
|
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
63 |
|
63572 | 64 |
text \<open>This form avoids giant explosions in proofs. NOTE USE OF \<open>\<equiv>\<close>.\<close> |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
65 |
lemma def_wfrec: "f \<equiv> wfrec R F \<Longrightarrow> wf R \<Longrightarrow> f a = F (cut f R a) a" |
63572 | 66 |
by (auto intro: wfrec) |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
67 |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
68 |
|
60758 | 69 |
subsubsection \<open>Well-founded recursion via genuine fixpoints\<close> |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
70 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
71 |
lemma wfrec_fixpoint: |
63572 | 72 |
assumes wf: "wf R" |
73 |
and adm: "adm_wf R F" |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
74 |
shows "wfrec R F = F (wfrec R F)" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
75 |
proof (rule ext) |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
76 |
fix x |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
77 |
have "wfrec R F x = F (cut (wfrec R F) R x) x" |
63572 | 78 |
using wfrec[of R F] wf by simp |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
79 |
also |
63572 | 80 |
have "\<And>y. (y, x) \<in> R \<Longrightarrow> cut (wfrec R F) R x y = wfrec R F y" |
81 |
by (auto simp add: cut_apply) |
|
82 |
then have "F (cut (wfrec R F) R x) x = F (wfrec R F) x" |
|
83 |
using adm adm_wf_def[of R F] by auto |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
84 |
finally show "wfrec R F x = F (wfrec R F) x" . |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
85 |
qed |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
86 |
|
74749
329cb9e6b184
A tiny bit of tidying connected with Zorn's Lemma
paulson <lp15@cam.ac.uk>
parents:
71544
diff
changeset
|
87 |
lemma wfrec_def_adm: "f \<equiv> wfrec R F \<Longrightarrow> wf R \<Longrightarrow> adm_wf R F \<Longrightarrow> f = F f" |
329cb9e6b184
A tiny bit of tidying connected with Zorn's Lemma
paulson <lp15@cam.ac.uk>
parents:
71544
diff
changeset
|
88 |
using wfrec_fixpoint by simp |
329cb9e6b184
A tiny bit of tidying connected with Zorn's Lemma
paulson <lp15@cam.ac.uk>
parents:
71544
diff
changeset
|
89 |
|
63572 | 90 |
|
61799 | 91 |
subsection \<open>Wellfoundedness of \<open>same_fst\<close>\<close> |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
92 |
|
63572 | 93 |
definition same_fst :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> ('b \<times> 'b) set) \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set" |
94 |
where "same_fst P R = {((x', y'), (x, y)) . x' = x \<and> P x \<and> (y',y) \<in> R x}" |
|
69593 | 95 |
\<comment> \<open>For \<^const>\<open>wfrec\<close> declarations where the first n parameters |
60758 | 96 |
stay unchanged in the recursive call.\<close> |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
97 |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
98 |
lemma same_fstI [intro!]: "P x \<Longrightarrow> (y', y) \<in> R x \<Longrightarrow> ((x, y'), (x, y)) \<in> same_fst P R" |
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
99 |
by (simp add: same_fst_def) |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
100 |
|
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
101 |
lemma wf_same_fst: |
71544 | 102 |
assumes "\<And>x. P x \<Longrightarrow> wf (R x)" |
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
55210
diff
changeset
|
103 |
shows "wf (same_fst P R)" |
75669
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
104 |
proof - |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
105 |
have "\<And>a b Q. \<forall>a b. (\<forall>x. P a \<and> (x, b) \<in> R a \<longrightarrow> Q (a, x)) \<longrightarrow> Q (a, b) \<Longrightarrow> Q (a, b)" |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
106 |
proof - |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
107 |
fix Q a b |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
108 |
assume *: "\<forall>a b. (\<forall>x. P a \<and> (x,b) \<in> R a \<longrightarrow> Q (a,x)) \<longrightarrow> Q (a,b)" |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
109 |
show "Q(a,b)" |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
110 |
proof (cases "wf (R a)") |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
111 |
case True |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
112 |
then show ?thesis |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
113 |
by (induction b rule: wf_induct_rule) (use * in blast) |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
114 |
qed (use * assms in blast) |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
115 |
qed |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
116 |
then show ?thesis |
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
Fabian Huch <huch@in.tum.de>
parents:
74749
diff
changeset
|
117 |
by (clarsimp simp add: wf_def same_fst_def) |
71544 | 118 |
qed |
44014
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
119 |
|
88bd7d74a2c1
moved recursion combinator to HOL/Library/Wfrec.thy -- it is so fundamental and well-known that it should survive recdef
krauss
parents:
diff
changeset
|
120 |
end |