diff -r fb63942e470e -r 7195acc2fe93 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Fri Aug 05 16:36:03 2016 +0200 +++ b/src/HOL/Wellfounded.thy Fri Aug 05 18:14:28 2016 +0200 @@ -187,9 +187,7 @@ text \Well-foundedness of subsets\ lemma wf_subset: "wf r \ p \ r \ wf p" - apply (simp add: wf_eq_minimal) - apply fast - done + by (simp add: wf_eq_minimal) fast lemmas wfP_subset = wf_subset [to_pred] @@ -198,11 +196,12 @@ lemma wf_empty [iff]: "wf {}" by (simp add: wf_def) -lemma wfP_empty [iff]: - "wfP (\x y. False)" +lemma wfP_empty [iff]: "wfP (\x y. False)" proof - - have "wfP bot" by (fact wf_empty [to_pred bot_empty_eq2]) - then show ?thesis by (simp add: bot_fun_def) + have "wfP bot" + by (fact wf_empty [to_pred bot_empty_eq2]) + then show ?thesis + by (simp add: bot_fun_def) qed lemma wf_Int1: "wf r \ wf (r \ r')" @@ -217,9 +216,10 @@ shows "wf R" proof (rule wfI_pf) fix A assume "A \ R `` A" - then have "A \ (R ^^ n) `` A" by (induct n) force+ - with \wf (R ^^ n)\ - show "A = {}" by (rule wfE_pf) + then have "A \ (R ^^ n) `` A" + by (induct n) force+ + with \wf (R ^^ n)\ show "A = {}" + by (rule wfE_pf) qed text \Well-foundedness of \insert\.\ @@ -257,7 +257,7 @@ apply (case_tac "\p. f p \ Q") apply (erule_tac x = "{p. f p \ Q}" in allE) apply (fast dest: inj_onD) -apply blast + apply blast done @@ -379,7 +379,7 @@ qed qed -lemma wf_comp_self: "wf R = wf (R O R)" \ \special case\ +lemma wf_comp_self: "wf R \ wf (R O R)" \ \special case\ by (rule wf_union_merge [where S = "{}", simplified]) @@ -390,12 +390,13 @@ lemma qc_wf_relto_iff: assumes "R O S \ (R \ S)\<^sup>* O R" \ \R quasi-commutes over S\ shows "wf (S\<^sup>* O R O S\<^sup>*) \ wf R" - (is "wf ?S \ _") + (is "wf ?S \ _") proof show "wf R" if "wf ?S" proof - have "R \ ?S" by auto - with that show "wf R" using wf_subset by auto + with wf_subset [of ?S] that show "wf R" + by auto qed next show "wf ?S" if "wf R" @@ -607,10 +608,7 @@ qed lemma accp_downwards_aux: "r\<^sup>*\<^sup>* b a \ accp r a \ accp r b" - apply (erule rtranclp_induct) - apply blast - apply (blast dest: accp_downward) - done + by (erule rtranclp_induct) (blast dest: accp_downward)+ theorem accp_downwards: "accp r a \ r\<^sup>*\<^sup>* b a \ accp r b" by (blast dest: accp_downwards_aux) @@ -691,7 +689,7 @@ text \Inverse Image\ lemma wf_inv_image [simp,intro!]: "wf r \ wf (inv_image r f)" - for f :: "'a \ 'b" + for f :: "'a \ 'b" apply (simp add: inv_image_def wf_eq_minimal) apply clarify apply (subgoal_tac "\w::'b. w \ {w. \x::'a. x \ Q \ f x = w}") @@ -776,7 +774,7 @@ done lemma trans_finite_psubset: "trans finite_psubset" - by (auto simp add: finite_psubset_def less_le trans_def) + by (auto simp: finite_psubset_def less_le trans_def) lemma in_finite_psubset[simp]: "(A, B) \ finite_psubset \ A \ B \ finite B" unfolding finite_psubset_def by auto