# HG changeset patch # User haftmann # Date 1291736034 -3600 # Node ID 36cec0e3491fbfb0f3262aa8e386336e538e2119 # Parent 6fabc0414055cde8cb53a6a1bfd4249f4243e8dc more concise case names; proved extensionality diff -r 6fabc0414055 -r 36cec0e3491f src/HOL/Quotient_Examples/FSet.thy --- a/src/HOL/Quotient_Examples/FSet.thy Tue Dec 07 16:33:54 2010 +0100 +++ b/src/HOL/Quotient_Examples/FSet.thy Tue Dec 07 16:33:54 2010 +0100 @@ -958,28 +958,28 @@ section {* Induction and Cases rules for fsets *} -lemma fset_exhaust [case_names empty_fset insert_fset, cases type: fset]: +lemma fset_exhaust [case_names empty insert, cases type: fset]: assumes empty_fset_case: "S = {||} \ P" and insert_fset_case: "\x S'. S = insert_fset x S' \ P" shows "P" using assms by (lifting list.exhaust) -lemma fset_induct [case_names empty_fset insert_fset]: +lemma fset_induct [case_names empty insert]: assumes empty_fset_case: "P {||}" and insert_fset_case: "\x S. P S \ P (insert_fset x S)" shows "P S" using assms by (descending) (blast intro: list.induct) -lemma fset_induct_stronger [case_names empty_fset insert_fset, induct type: fset]: +lemma fset_induct_stronger [case_names empty insert, induct type: fset]: assumes empty_fset_case: "P {||}" and insert_fset_case: "\x S. \x |\| S; P S\ \ P (insert_fset x S)" shows "P S" proof(induct S rule: fset_induct) - case empty_fset + case empty show "P {||}" using empty_fset_case by simp next - case (insert_fset x S) + case (insert x S) have "P S" by fact then show "P (insert_fset x S)" using insert_fset_case by (cases "x |\| S") (simp_all) @@ -990,10 +990,10 @@ and card_fset_Suc_case: "\S T. Suc (card_fset S) = (card_fset T) \ P S \ P T" shows "P S" proof (induct S) - case empty_fset + case empty show "P {||}" by (rule empty_fset_case) next - case (insert_fset x S) + case (insert x S) have h: "P S" by fact have "x |\| S" by fact then have "Suc (card_fset S) = card_fset (insert_fset x S)" @@ -1058,7 +1058,22 @@ apply simp_all done +text {* Extensionality *} +lemma fset_eqI: + assumes "\x. x \ fset A \ x \ fset B" + shows "A = B" +using assms proof (induct A arbitrary: B) + case empty then show ?case + by (auto simp add: in_fset none_in_empty_fset [symmetric] sym) +next + case (insert x A) + from insert.prems insert.hyps(1) have "\z. z \ fset A \ z \ fset (B - {|x|})" + by (auto simp add: in_fset) + then have "A = B - {|x|}" by (rule insert.hyps(2)) + moreover with insert.prems [symmetric, of x] have "x |\| B" by (simp add: in_fset) + ultimately show ?case by (metis in_fset_mdef) +qed subsection {* alternate formulation with a different decomposition principle and a proof of equivalence *}