# HG changeset patch # User huffman # Date 1269989210 25200 # Node ID b846881928ea8d3eb894fefefe0d5d0b40ca057f # Parent 10563d88c0b6ae07a8ceebacca3b83b72079f254 simplify fold_graph proofs diff -r 10563d88c0b6 -r b846881928ea src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Tue Mar 30 23:12:55 2010 +0200 +++ b/src/HOL/Finite_Set.thy Tue Mar 30 15:46:50 2010 -0700 @@ -616,125 +616,56 @@ subsubsection{*From @{const fold_graph} to @{term fold}*} -lemma image_less_Suc: "h ` {i. i < Suc m} = insert (h m) (h ` {i. i < m})" - by (auto simp add: less_Suc_eq) - -lemma insert_image_inj_on_eq: - "[|insert (h m) A = h ` {i. i < Suc m}; h m \ A; - inj_on h {i. i < Suc m}|] - ==> A = h ` {i. i < m}" -apply (auto simp add: image_less_Suc inj_on_def) -apply (blast intro: less_trans) -done - -lemma insert_inj_onE: - assumes aA: "insert a A = h`{i::nat. i A" - and inj_on: "inj_on h {i::nat. ihm m. inj_on hm {i::nat. i A" by (simp add: swap_def hkeq anot) - show "insert (?hm m) A = ?hm ` {i. i < Suc m}" - using aA hkeq nSuc klessn - by (auto simp add: swap_def image_less_Suc fun_upd_image - less_Suc_eq inj_on_image_set_diff [OF inj_on]) - qed - qed -qed - context fun_left_comm begin -lemma fold_graph_determ_aux: - "A = h`{i::nat. i inj_on h {i. i fold_graph f z A x \ fold_graph f z A x' - \ x' = x" -proof (induct n arbitrary: A x x' h rule: less_induct) - case (less n) - have IH: "\m h A x x'. m < n \ A = h ` {i. i inj_on h {i. i fold_graph f z A x - \ fold_graph f z A x' \ x' = x" by fact - have Afoldx: "fold_graph f z A x" and Afoldx': "fold_graph f z A x'" - and A: "A = h`{i. i a \ A \ \y'. y = f a y' \ fold_graph f z (A - {a}) y'" +proof (induct set: fold_graph) + case (insertI x A y) show ?case + proof (cases "x = a") + assume "x = a" with insertI show ?case by auto next - fix B b u - assume AbB: "A = insert b B" and x: "x = f b u" - and notinB: "b \ B" and Bu: "fold_graph f z B u" - show "x'=x" - proof (rule fold_graph.cases [OF Afoldx']) - assume "A = {}" and "x' = z" - with AbB show "x' = x" by blast - next - fix C c v - assume AcC: "A = insert c C" and x': "x' = f c v" - and notinC: "c \ C" and Cv: "fold_graph f z C v" - from A AbB have Beq: "insert b B = h`{i. i c" - let ?D = "B - {c}" - have B: "B = insert c ?D" and C: "C = insert b ?D" - using AbB AcC notinB notinC diff by(blast elim!:equalityE)+ - have "finite A" by(rule fold_graph_imp_finite [OF Afoldx]) - with AbB have "finite ?D" by simp - then obtain d where Dfoldd: "fold_graph f z ?D d" - using finite_imp_fold_graph by iprover - moreover have cinB: "c \ B" using B by auto - ultimately have "fold_graph f z B (f c d)" by(rule Diff1_fold_graph) - hence "f c d = u" by (rule IH [OF lessB Beq inj_onB Bu]) - moreover have "f b d = v" - proof (rule IH[OF lessC Ceq inj_onC Cv]) - show "fold_graph f z C (f b d)" using C notinB Dfoldd by fastsimp - qed - ultimately show ?thesis - using fun_left_comm [of c b] x x' by (auto simp add: o_def) - qed - qed + assume "x \ a" + then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'" + using insertI by auto + have 1: "f x y = f a (f x y')" + unfolding y by (rule fun_left_comm) + have 2: "fold_graph f z (insert x A - {a}) (f x y')" + using y' and `x \ a` and `x \ A` + by (simp add: insert_Diff_if fold_graph.insertI) + from 1 2 show ?case by fast qed -qed +qed simp + +lemma fold_graph_insertE: + assumes "fold_graph f z (insert x A) v" and "x \ A" + obtains y where "v = f x y" and "fold_graph f z A y" +using assms by (auto dest: fold_graph_insertE_aux [OF _ insertI1]) lemma fold_graph_determ: "fold_graph f z A x \ fold_graph f z A y \ y = x" -apply (frule fold_graph_imp_finite [THEN finite_imp_nat_seg_image_inj_on]) -apply (blast intro: fold_graph_determ_aux [rule_format]) -done +proof (induct arbitrary: y set: fold_graph) + case (insertI x A y v) + from `fold_graph f z (insert x A) v` and `x \ A` + obtain y' where "v = f x y'" and "fold_graph f z A y'" + by (rule fold_graph_insertE) + from `fold_graph f z A y'` have "y' = y" by (rule insertI) + with `v = f x y'` show "v = f x y" by simp +qed fast lemma fold_equality: "fold_graph f z A y \ fold f z A = y" by (unfold fold_def) (blast intro: fold_graph_determ) +lemma fold_graph_fold: "finite A \ fold_graph f z A (fold f z A)" +unfolding fold_def +apply (rule theI') +apply (rule ex_ex1I) +apply (erule finite_imp_fold_graph) +apply (erule (1) fold_graph_determ) +done + text{* The base case for @{text fold}: *} lemma (in -) fold_empty [simp]: "fold f z {} = z" @@ -742,21 +673,11 @@ text{* The various recursion equations for @{const fold}: *} -lemma fold_insert_aux: "x \ A - \ fold_graph f z (insert x A) v \ - (\y. fold_graph f z A y \ v = f x y)" -apply auto -apply (rule_tac A1 = A and f1 = f in finite_imp_fold_graph [THEN exE]) - apply (fastsimp dest: fold_graph_imp_finite) -apply (blast intro: fold_graph_determ) -done - lemma fold_insert [simp]: "finite A ==> x \ A ==> fold f z (insert x A) = f x (fold f z A)" -apply (simp add: fold_def fold_insert_aux) -apply (rule the_equality) - apply (auto intro: finite_imp_fold_graph - cong add: conj_cong simp add: fold_def[symmetric] fold_equality) +apply (rule fold_equality) +apply (erule fold_graph.insertI) +apply (erule fold_graph_fold) done lemma fold_fun_comm: @@ -1185,7 +1106,7 @@ interpret fun_left_comm "op *::'a \ 'a \ 'a" by (rule fun_left_comm) from assms show ?thesis proof (induct rule: fold_graph.induct) - case emptyI thus ?case by (force simp add: fold_insert_aux mult_commute) + case emptyI show ?case by (subst mult_commute [of z b], fast) next case (insertI x A y) have "fold_graph times z (insert x (insert b A)) (x * (z * y))"