src/HOL/BNF_Fixpoint_Base.thy
changeset 62905 52c5a25e0c96
parent 62335 e85c42f4f30a
child 63046 8053ef5a0174
equal deleted inserted replaced
62904:94535e6dd168 62905:52c5a25e0c96
   240   by (simp only: eq_onp_same_args)
   240   by (simp only: eq_onp_same_args)
   241 
   241 
   242 lemma top_conj: "top x \<and> P \<longleftrightarrow> P" "P \<and> top x \<longleftrightarrow> P"
   242 lemma top_conj: "top x \<and> P \<longleftrightarrow> P" "P \<and> top x \<longleftrightarrow> P"
   243   by blast+
   243   by blast+
   244 
   244 
       
   245 lemma fst_convol': "fst (\<langle>f, g\<rangle> x) = f x"
       
   246   using fst_convol unfolding convol_def by simp
       
   247 
       
   248 lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x"
       
   249   using snd_convol unfolding convol_def by simp
       
   250 
       
   251 lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f"
       
   252   unfolding convol_def by auto
       
   253 
       
   254 lemma convol_expand_snd':
       
   255   assumes "(fst o f = g)"
       
   256   shows "h = snd o f \<longleftrightarrow> \<langle>g, h\<rangle> = f"
       
   257 proof -
       
   258   from assms have *: "\<langle>g, snd o f\<rangle> = f" by (rule convol_expand_snd)
       
   259   then have "h = snd o f \<longleftrightarrow> h = snd o \<langle>g, snd o f\<rangle>" by simp
       
   260   moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol)
       
   261   moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
       
   262   ultimately show ?thesis by simp
       
   263 qed
       
   264 
       
   265 lemma case_sum_expand_Inr_pointfree: "f o Inl = g \<Longrightarrow> case_sum g (f o Inr) = f"
       
   266   by (auto split: sum.splits)
       
   267 
       
   268 lemma case_sum_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> case_sum g h = f"
       
   269   by (rule iffI) (auto simp add: fun_eq_iff split: sum.splits)
       
   270 
       
   271 lemma case_sum_expand_Inr: "f o Inl = g \<Longrightarrow> f x = case_sum g (f o Inr) x"
       
   272   by (auto split: sum.splits)
       
   273 
       
   274 lemma id_transfer: "rel_fun A A id id"
       
   275   unfolding rel_fun_def by simp
       
   276 
       
   277 lemma fst_transfer: "rel_fun (rel_prod A B) A fst fst"
       
   278   unfolding rel_fun_def by simp
       
   279 
       
   280 lemma snd_transfer: "rel_fun (rel_prod A B) B snd snd"
       
   281   unfolding rel_fun_def by simp
       
   282 
       
   283 lemma convol_transfer:
       
   284   "rel_fun (rel_fun R S) (rel_fun (rel_fun R T) (rel_fun R (rel_prod S T))) BNF_Def.convol BNF_Def.convol"
       
   285   unfolding rel_fun_def convol_def by auto
       
   286 
       
   287 ML_file "Tools/BNF/bnf_fp_util_tactics.ML"
   245 ML_file "Tools/BNF/bnf_fp_util.ML"
   288 ML_file "Tools/BNF/bnf_fp_util.ML"
   246 ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML"
   289 ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML"
   247 ML_file "Tools/BNF/bnf_fp_def_sugar.ML"
   290 ML_file "Tools/BNF/bnf_fp_def_sugar.ML"
   248 ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML"
   291 ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML"
   249 ML_file "Tools/BNF/bnf_fp_n2m.ML"
   292 ML_file "Tools/BNF/bnf_fp_n2m.ML"