diff -r 4a15873c4ec9 -r 41ee3bfccb4d src/HOL/BNF/Examples/Derivation_Trees/Prelim.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/BNF/Examples/Derivation_Trees/Prelim.thy Tue Oct 16 13:09:46 2012 +0200 @@ -0,0 +1,67 @@ +(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Prelim.thy + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Preliminaries. +*) + +header {* Preliminaries *} + +theory Prelim +imports "../../BNF" +begin + +declare fset_to_fset[simp] + +lemma fst_snd_convol_o[simp]: " = s" +apply(rule ext) by (simp add: convol_def) + +abbreviation sm_abbrev (infix "\" 60) +where "f \ g \ Sum_Type.sum_map f g" + +lemma sum_map_InlD: "(f \ g) z = Inl x \ \y. z = Inl y \ f y = x" +by (cases z) auto + +lemma sum_map_InrD: "(f \ g) z = Inr x \ \y. z = Inr y \ g y = x" +by (cases z) auto + +abbreviation sum_case_abbrev ("[[_,_]]" 800) +where "[[f,g]] \ Sum_Type.sum_case f g" + +lemma inj_Inl[simp]: "inj Inl" unfolding inj_on_def by auto +lemma inj_Inr[simp]: "inj Inr" unfolding inj_on_def by auto + +lemma Inl_oplus_elim: +assumes "Inl tr \ (id \ f) ` tns" +shows "Inl tr \ tns" +using assms apply clarify by (case_tac x, auto) + +lemma Inl_oplus_iff[simp]: "Inl tr \ (id \ f) ` tns \ Inl tr \ tns" +using Inl_oplus_elim +by (metis id_def image_iff sum_map.simps(1)) + +lemma Inl_m_oplus[simp]: "Inl -` (id \ f) ` tns = Inl -` tns" +using Inl_oplus_iff unfolding vimage_def by auto + +lemma Inr_oplus_elim: +assumes "Inr tr \ (id \ f) ` tns" +shows "\ n. Inr n \ tns \ f n = tr" +using assms apply clarify by (case_tac x, auto) + +lemma Inr_oplus_iff[simp]: +"Inr tr \ (id \ f) ` tns \ (\ n. Inr n \ tns \ f n = tr)" +apply (rule iffI) + apply (metis Inr_oplus_elim) +by (metis image_iff sum_map.simps(2)) + +lemma Inr_m_oplus[simp]: "Inr -` (id \ f) ` tns = f ` (Inr -` tns)" +using Inr_oplus_iff unfolding vimage_def by auto + +lemma Inl_Inr_image_cong: +assumes "Inl -` A = Inl -` B" and "Inr -` A = Inr -` B" +shows "A = B" +apply safe using assms apply(case_tac x, auto) by(case_tac x, auto) + + + +end \ No newline at end of file