# HG changeset patch # User wenzelm # Date 1475342988 -7200 # Node ID 6f7db4f8df4c4495655bdc2f01354d4eb6566271 # Parent f8e556c8ad6ffb97712bbd5e8ea28352e3690b8b tuned proofs; diff -r f8e556c8ad6f -r 6f7db4f8df4c src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Sat Oct 01 17:38:14 2016 +0200 +++ b/src/HOL/Hilbert_Choice.thy Sat Oct 01 19:29:48 2016 +0200 @@ -466,30 +466,48 @@ text \A relation is wellfounded iff it has no infinite descending chain.\ -lemma wf_iff_no_infinite_down_chain: "wf r \ (\ (\f. \i. (f (Suc i), f i) \ r))" - apply (simp only: wf_eq_minimal) - apply (rule iffI) - apply (rule notI) - apply (erule exE) - apply (erule_tac x = "{w. \i. w = f i}" in allE) - apply blast - apply (erule contrapos_np) - apply simp - apply clarify - apply (subgoal_tac "\n. rec_nat x (\i y. SOME z. z \ Q \ (z, y) \ r) n \ Q") - apply (rule_tac x = "rec_nat x (\i y. SOME z. z \ Q \ (z, y) \ r)" in exI) - apply (rule allI) - apply simp - apply (rule someI2_ex) - apply blast - apply blast - apply (rule allI) - apply (induct_tac n) - apply simp_all - apply (rule someI2_ex) - apply blast - apply blast - done +lemma wf_iff_no_infinite_down_chain: "wf r \ (\f. \i. (f (Suc i), f i) \ r)" + (is "_ \ \ ?ex") +proof + assume "wf r" + show "\ ?ex" + proof + assume ?ex + then obtain f where f: "(f (Suc i), f i) \ r" for i + by blast + from \wf r\ have minimal: "x \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" for x Q + by (auto simp: wf_eq_minimal) + let ?Q = "{w. \i. w = f i}" + fix n + have "f n \ ?Q" by blast + from minimal [OF this] obtain j where "(y, f j) \ r \ y \ ?Q" for y by blast + with this [OF \(f (Suc j), f j) \ r\] have "f (Suc j) \ ?Q" by simp + then show False by blast + qed +next + assume "\ ?ex" + then show "wf r" + proof (rule contrapos_np) + assume "\ wf r" + then obtain Q x where x: "x \ Q" and rec: "z \ Q \ \y. (y, z) \ r \ y \ Q" for z + by (auto simp add: wf_eq_minimal) + obtain descend :: "nat \ 'a" + where descend_0: "descend 0 = x" + and descend_Suc: "descend (Suc n) = (SOME y. y \ Q \ (y, descend n) \ r)" for n + by (rule that [of "rec_nat x (\_ rec. (SOME y. y \ Q \ (y, rec) \ r))"]) simp_all + have descend_Q: "descend n \ Q" for n + proof (induct n) + case 0 + with x show ?case by (simp only: descend_0) + next + case Suc + then show ?case by (simp only: descend_Suc) (rule someI2_ex; use rec in blast) + qed + have "(descend (Suc i), descend i) \ r" for i + by (simp only: descend_Suc) (rule someI2_ex; use descend_Q rec in blast) + then show "\f. \i. (f (Suc i), f i) \ r" by blast + qed +qed lemma wf_no_infinite_down_chainE: assumes "wf r" diff -r f8e556c8ad6f -r 6f7db4f8df4c src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Sat Oct 01 17:38:14 2016 +0200 +++ b/src/HOL/Inductive.thy Sat Oct 01 19:29:48 2016 +0200 @@ -23,10 +23,10 @@ where "lfp f = Inf {u. f u \ u}" lemma lfp_lowerbound: "f A \ A \ lfp f \ A" - by (auto simp add: lfp_def intro: Inf_lower) + unfolding lfp_def by (rule Inf_lower) simp lemma lfp_greatest: "(\u. f u \ u \ A \ u) \ A \ lfp f" - by (auto simp add: lfp_def intro: Inf_greatest) + unfolding lfp_def by (rule Inf_greatest) simp end