# HG changeset patch # User Christian Urban # Date 1264463260 -3600 # Node ID 762f0eae88fdf033d7d3a37bfabbb9f8a667b994 # Parent 7781c7cbd27eca0d21e1e7f1e59c224d7a583c1d used split_option_all lemma diff -r 7781c7cbd27e -r 762f0eae88fd Quot/QuotOption.thy --- a/Quot/QuotOption.thy Tue Jan 26 00:18:48 2010 +0100 +++ b/Quot/QuotOption.thy Tue Jan 26 00:47:40 2010 +0100 @@ -17,19 +17,23 @@ declare [[map option = (Option.map, option_rel)]] +text {* should probably be in Option.thy *} +lemma split_option_all: + shows "(\x. P x) \ P None \ (\a. P (Some a))" +apply(auto) +apply(case_tac x) +apply(simp_all) +done + lemma option_quotient[quot_thm]: assumes q: "Quotient R Abs Rep" shows "Quotient (option_rel R) (Option.map Abs) (Option.map Rep)" unfolding Quotient_def - apply(auto) - apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) - apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) - apply(case_tac [!] r) - apply(case_tac [!] s) - apply(simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q] ) + apply(simp add: split_option_all) + apply(simp add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) using q unfolding Quotient_def - apply(blast)+ + apply(blast) done lemma option_equivp[quot_equiv]: @@ -37,16 +41,10 @@ shows "equivp (option_rel R)" apply(rule equivpI) unfolding reflp_def symp_def transp_def - apply(auto) - apply(case_tac [!] x) - apply(simp_all add: equivp_reflp[OF a]) - apply(case_tac [!] y) - apply(simp_all add: equivp_symp[OF a]) - apply(case_tac [!] z) - apply(simp_all) - apply(clarify) - apply(rule equivp_transp[OF a]) - apply(assumption)+ + apply(simp_all add: split_option_all) + apply(blast intro: equivp_reflp[OF a]) + apply(blast intro: equivp_symp[OF a]) + apply(blast intro: equivp_transp[OF a]) done lemma option_None_rsp[quot_respect]: @@ -74,18 +72,11 @@ lemma option_map_id[id_simps]: shows "Option.map id \ id" apply (rule eq_reflection) - apply (auto simp add: expand_fun_eq) - apply (case_tac x) - apply (auto) - done + by (simp add: expand_fun_eq split_option_all) lemma option_rel_eq[id_simps]: shows "option_rel (op =) \ (op =)" apply(rule eq_reflection) - apply(auto simp add: expand_fun_eq) - apply(case_tac x) - apply(case_tac [!] xa) - apply(auto) - done + by (simp add: expand_fun_eq split_option_all) end