diff -r 5004aca20821 -r 882091eb1e9a src/HOL/BNF_Def.thy --- a/src/HOL/BNF_Def.thy Fri Jun 27 10:11:44 2014 +0200 +++ b/src/HOL/BNF_Def.thy Fri Jun 27 10:11:44 2014 +0200 @@ -1,5 +1,6 @@ (* Title: HOL/BNF_Def.thy Author: Dmitriy Traytel, TU Muenchen + Author: Jasmin Blanchette, TU Muenchen Copyright 2012 Definition of bounded natural functors. @@ -8,12 +9,47 @@ header {* Definition of Bounded Natural Functors *} theory BNF_Def -imports BNF_Util Fun_Def_Base +imports BNF_Cardinal_Arithmetic Fun_Def_Base keywords "print_bnfs" :: diag and "bnf" :: thy_goal begin +definition + rel_fun :: "('a \ 'c \ bool) \ ('b \ 'd \ bool) \ ('a \ 'b) \ ('c \ 'd) \ bool" +where + "rel_fun A B = (\f g. \x y. A x y \ B (f x) (g y))" + +lemma rel_funI [intro]: + assumes "\x y. A x y \ B (f x) (g y)" + shows "rel_fun A B f g" + using assms by (simp add: rel_fun_def) + +lemma rel_funD: + assumes "rel_fun A B f g" and "A x y" + shows "B (f x) (g y)" + using assms by (simp add: rel_fun_def) + +definition collect where +"collect F x = (\f \ F. f x)" + +lemma fstI: "x = (y, z) \ fst x = y" +by simp + +lemma sndI: "x = (y, z) \ snd x = z" +by simp + +lemma bijI': "\\x y. (f x = f y) = (x = y); \y. \x. y = f x\ \ bij f" +unfolding bij_def inj_on_def by auto blast + +(* Operator: *) +definition "Gr A f = {(a, f a) | a. a \ A}" + +definition "Grp A f = (\a b. b = f a \ a \ A)" + +definition vimage2p where + "vimage2p f g R = (\x y. R (f x) (g y))" + lemma collect_comp: "collect F \ g = collect ((\f. f \ g) ` F)" by (rule ext) (auto simp only: comp_apply collect_def) @@ -152,6 +188,8 @@ lemma vimage2p_Grp: "vimage2p f g P = Grp UNIV f OO P OO (Grp UNIV g)\\" unfolding vimage2p_def Grp_def by auto +ML_file "Tools/BNF/bnf_util.ML" +ML_file "Tools/BNF/bnf_tactics.ML" ML_file "Tools/BNF/bnf_def_tactics.ML" ML_file "Tools/BNF/bnf_def.ML"