# HG changeset patch # User desharna # Date 1711536582 -3600 # Node ID 38803a6b3357dbabd1a97005d5941e79653ede30 # Parent b0a46cf73aa478c6559cc8d92276664e78420bf2 added lemma wfp_on_image and author name to theory diff -r b0a46cf73aa4 -r 38803a6b3357 NEWS --- a/NEWS Wed Mar 27 10:54:47 2024 +0100 +++ b/NEWS Wed Mar 27 11:49:42 2024 +0100 @@ -120,7 +120,6 @@ predicates such as asymp, transp, or totalp. - Added lemmas. wellorder.wfp_on_less[simp] - wfp_iff_ex_minimal wf_iff_ex_minimal wf_onE_pf wf_onI_pf @@ -131,10 +130,12 @@ wf_on_iff_wf wf_on_induct wf_on_subset + wfp_iff_ex_minimal wfp_on_antimono wfp_on_antimono_strong wfp_on_if_convertible_to_wfp_on wfp_on_iff_ex_minimal + wfp_on_image wfp_on_induct wfp_on_inv_imagep wfp_on_subset diff -r b0a46cf73aa4 -r 38803a6b3357 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Wed Mar 27 10:54:47 2024 +0100 +++ b/src/HOL/Wellfounded.thy Wed Mar 27 11:49:42 2024 +0100 @@ -4,6 +4,7 @@ Author: Konrad Slind Author: Alexander Krauss Author: Andrei Popescu, TU Muenchen + Author: Martin Desharnais, MPI-INF Saarbruecken *) section \Well-founded Recursion\ @@ -490,6 +491,61 @@ lemma wf_map_prod_image: "wf r \ inj f \ wf (map_prod f f ` r)" by(rule wf_map_prod_image_Dom_Ran) (auto dest: inj_onD) +lemma wfp_on_image: "wfp_on (f ` A) R \ wfp_on A (\a b. R (f a) (f b))" +proof (rule iffI) + assume hyp: "wfp_on (f ` A) R" + show "wfp_on A (\a b. R (f a) (f b))" + unfolding wfp_on_iff_ex_minimal + proof (intro allI impI) + fix B + assume "B \ A" and "B \ {}" + hence "f ` B \ f ` A" and "f ` B \ {}" + unfolding atomize_conj image_is_empty + using image_mono by iprover + hence "\z\f ` B. \y. R y z \ y \ f ` B" + using hyp[unfolded wfp_on_iff_ex_minimal, rule_format] by iprover + then obtain fz where "fz \ f ` B" and fz_max: "\y. R y fz \ y \ f ` B" .. + + obtain z where "z \ B" and "fz = f z" + using \fz \ f ` B\ unfolding image_iff .. + + show "\z\B. \y. R (f y) (f z) \ y \ B" + proof (intro bexI allI impI) + show "z \ B" + using \z \ B\ . + next + fix y assume "R (f y) (f z)" + hence "f y \ f ` B" + using fz_max \fz = f z\ by iprover + thus "y \ B" + by (rule contrapos_nn) (rule imageI) + qed + qed +next + assume hyp: "wfp_on A (\a b. R (f a) (f b))" + show "wfp_on (f ` A) R" + unfolding wfp_on_iff_ex_minimal + proof (intro allI impI) + fix fA + assume "fA \ f ` A" and "fA \ {}" + then obtain A' where "A' \ A" and "A' \ {}" and "fA = f ` A'" + by (auto simp only: subset_image_iff) + + obtain z where "z \ A'" and z_max: "\y. R (f y) (f z) \ y \ A'" + using hyp[unfolded wfp_on_iff_ex_minimal, rule_format, OF \A' \ A\ \A' \ {}\] by blast + + show "\z\fA. \y. R y z \ y \ fA" + proof (intro bexI allI impI) + show "f z \ fA" + unfolding \fA = f ` A'\ + using imageI[OF \z \ A'\] . + next + show "\y. R y (f z) \ y \ fA" + unfolding \fA = f ` A'\ + using z_max by auto + qed + qed +qed subsection \Well-Foundedness Results for Unions\