# HG changeset patch # User paulson # Date 1531862307 -3600 # Node ID 7dc9fe795daea31b7d817f9a76f5877fec8c621d # Parent 5e15795788d3ceff9527c0b24aa3d61df14ff125 more de-applying diff -r 5e15795788d3 -r 7dc9fe795dae src/HOL/List.thy --- a/src/HOL/List.thy Mon Jul 16 23:33:38 2018 +0100 +++ b/src/HOL/List.thy Tue Jul 17 22:18:27 2018 +0100 @@ -5804,14 +5804,11 @@ by (induct n arbitrary: xs ys) auto lemma wf_lex [intro!]: "wf r ==> wf (lex r)" -apply (unfold lex_def) -apply (rule wf_UN) -apply (blast intro: wf_lexn, clarify) -apply (rename_tac m n) -apply (subgoal_tac "m \ n") - prefer 2 apply blast -apply (blast dest: lexn_length not_sym) -done + apply (unfold lex_def) + apply (rule wf_UN) + apply (simp add: wf_lexn) + apply (metis DomainE Int_emptyI RangeE lexn_length) + done lemma lexn_conv: "lexn r n = diff -r 5e15795788d3 -r 7dc9fe795dae src/HOL/MicroJava/DFA/Listn.thy --- a/src/HOL/MicroJava/DFA/Listn.thy Mon Jul 16 23:33:38 2018 +0100 +++ b/src/HOL/MicroJava/DFA/Listn.thy Tue Jul 17 22:18:27 2018 +0100 @@ -323,12 +323,10 @@ apply (blast intro: lesssub_list_impl_same_size) apply (rule wf_UN) prefer 2 - apply clarify apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) -apply clarify apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) @@ -353,7 +351,7 @@ apply blast apply clarify apply (thin_tac "m \ M") -apply (thin_tac "maxA#xs \ M") + apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption diff -r 5e15795788d3 -r 7dc9fe795dae src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Mon Jul 16 23:33:38 2018 +0100 +++ b/src/HOL/Wellfounded.thy Tue Jul 17 22:18:27 2018 +0100 @@ -132,13 +132,9 @@ qed lemma wf_eq_minimal: "wf r \ (\Q x. x \ Q \ (\z\Q. \y. (y, z) \ r \ y \ Q))" - apply auto - apply (erule wfE_min) - apply assumption - apply blast - apply (rule wfI_min) - apply auto - done + apply (rule iffI) + apply (blast intro: elim!: wfE_min) + by (rule wfI_min) auto lemmas wfP_eq_minimal = wf_eq_minimal [to_pred] @@ -223,30 +219,45 @@ qed text \Well-foundedness of \insert\.\ -lemma wf_insert [iff]: "wf (insert (y, x) r) \ wf r \ (x, y) \ r\<^sup>*" - apply (rule iffI) - apply (blast elim: wf_trancl [THEN wf_irrefl] - intro: rtrancl_into_trancl1 wf_subset rtrancl_mono [THEN [2] rev_subsetD]) - apply (simp add: wf_eq_minimal) - apply safe - apply (rule allE) - apply assumption - apply (erule impE) - apply blast - apply (erule bexE) - apply (rename_tac a, case_tac "a = x") - prefer 2 - apply blast - apply (case_tac "y \ Q") - prefer 2 - apply blast - apply (rule_tac x = "{z. z \ Q \ (z,y) \ r\<^sup>*}" in allE) - apply assumption - apply (erule_tac V = "\Q. (\x. x \ Q) \ P Q" for P in thin_rl) - (*essential for speed*) - (*blast with new substOccur fails*) - apply (fast intro: converse_rtrancl_into_rtrancl) - done +lemma wf_insert [iff]: "wf (insert (y,x) r) \ wf r \ (x,y) \ r\<^sup>*" (is "?lhs = ?rhs") +proof + assume ?lhs then show ?rhs + by (blast elim: wf_trancl [THEN wf_irrefl] + intro: rtrancl_into_trancl1 wf_subset rtrancl_mono [THEN subsetD]) +next + assume R: ?rhs + then have R': "Q \ {} \ (\z\Q. \y. (y, z) \ r \ y \ Q)" for Q + by (auto simp: wf_eq_minimal) + show ?lhs + unfolding wf_eq_minimal + proof clarify + fix Q :: "'a set" and q + assume "q \ Q" + then obtain a where "a \ Q" and a: "\y. (y, a) \ r \ y \ Q" + using R by (auto simp: wf_eq_minimal) + show "\z\Q. \y'. (y', z) \ insert (y, x) r \ y' \ Q" + proof (cases "a=x") + case True + show ?thesis + proof (cases "y \ Q") + case True + then obtain z where "z \ Q" "(z, y) \ r\<^sup>*" + "\z'. (z', z) \ r \ z' \ Q \ (z', y) \ r\<^sup>*" + using R' [of "{z \ Q. (z,y) \ r\<^sup>*}"] by auto + with R show ?thesis + by (rule_tac x="z" in bexI) (blast intro: rtrancl_trans) + next + case False + then show ?thesis + using a \a \ Q\ by blast + qed + next + case False + with a \a \ Q\ show ?thesis + by blast + qed + qed +qed subsubsection \Well-foundedness of image\ @@ -307,20 +318,29 @@ text \Well-foundedness of indexed union with disjoint domains and ranges.\ lemma wf_UN: - assumes "\i\I. wf (r i)" - and "\i\I. \j\I. r i \ r j \ Domain (r i) \ Range (r j) = {}" + assumes r: "\i. i \ I \ wf (r i)" + and disj: "\i j. \i \ I; j \ I; r i \ r j\ \ Domain (r i) \ Range (r j) = {}" shows "wf (\i\I. r i)" - using assms - apply (simp only: wf_eq_minimal) - apply clarify - apply (rename_tac A a, case_tac "\i\I. \a\A. \b\A. (b, a) \ r i") - prefer 2 - apply force - apply clarify - apply (drule bspec, assumption) - apply (erule_tac x="{a. a \ A \ (\b\A. (b, a) \ r i) }" in allE) - apply (blast elim!: allE) - done + unfolding wf_eq_minimal +proof clarify + fix A and a :: "'b" + assume "a \ A" + show "\z\A. \y. (y, z) \ UNION I r \ y \ A" + proof (cases "\i\I. \a\A. \b\A. (b, a) \ r i") + case True + then obtain i b c where ibc: "i \ I" "b \ A" "c \ A" "(c,b) \ r i" + by blast + have ri: "\Q. Q \ {} \ \z\Q. \y. (y, z) \ r i \ y \ Q" + using r [OF \i \ I\] unfolding wf_eq_minimal by auto + show ?thesis + using ri [of "{a. a \ A \ (\b\A. (b, a) \ r i) }"] ibc disj + by blast + next + case False + with \a \ A\ show ?thesis + by blast + qed +qed lemma wfP_SUP: "\i. wfP (r i) \ \i j. r i \ r j \ inf (Domainp (r i)) (Rangep (r j)) = bot \ @@ -482,11 +502,14 @@ subsubsection \Wellfoundedness of finite acyclic relations\ -lemma finite_acyclic_wf [rule_format]: "finite r \ acyclic r \ wf r" - apply (erule finite_induct) - apply blast - apply (simp add: split_tupled_all) - done +lemma finite_acyclic_wf: + assumes "finite r" "acyclic r" shows "wf r" + using assms +proof (induction r rule: finite_induct) + case (insert x r) + then show ?case + by (cases x) simp +qed simp lemma finite_acyclic_wf_converse: "finite r \ acyclic r \ wf (r\)" apply (erule finite_converse [THEN iffD2, THEN finite_acyclic_wf]) @@ -595,8 +618,7 @@ apply (rule major [THEN accp.induct]) apply (rule hyp) apply (rule accp.accI) - apply fast - apply fast + apply auto done lemmas accp_induct_rule = accp_induct [rule_format, induct set: accp] @@ -631,8 +653,7 @@ theorem accp_wfPI: "\x. accp r x \ wfP r" apply (rule wfPUNIVI) apply (rule_tac P = P in accp_induct) - apply blast - apply blast + apply blast+ done theorem accp_wfPD: "wfP r \ accp r x" @@ -728,11 +749,8 @@ lemma wf_if_measure: "(\x. P x \ f(g x) < f x) \ wf {(y,x). P x \ y = g x}" for f :: "'a \ nat" - apply (insert wf_measure[of f]) - apply (simp only: measure_def inv_image_def less_than_def less_eq) - apply (erule wf_subset) - apply auto - done + using wf_measure[of f] unfolding measure_def inv_image_def less_than_def less_eq + by (rule wf_subset) auto subsubsection \Lexicographic combinations\ @@ -742,7 +760,7 @@ where "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \ ra \ a = a' \ (b, b') \ rb}" lemma wf_lex_prod [intro!]: "wf ra \ wf rb \ wf (ra <*lex*> rb)" - apply (unfold wf_def lex_prod_def) + unfolding wf_def lex_prod_def apply (rule allI) apply (rule impI) apply (simp only: split_paired_All)