| author | wenzelm | 
| Tue, 02 Jul 2024 23:29:46 +0200 | |
| changeset 80482 | 2136ecf06a4c | 
| parent 74886 | fa5476c54731 | 
| child 80932 | 261cd8722677 | 
| permissions | -rw-r--r-- | 
| 55059 | 1 | (* Title: HOL/BNF_Def.thy | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 57398 | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 57698 | 4 | Copyright 2012, 2013, 2014 | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 5 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 6 | Definition of bounded natural functors. | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 7 | *) | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 8 | |
| 60758 | 9 | section \<open>Definition of Bounded Natural Functors\<close> | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 10 | |
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 11 | theory BNF_Def | 
| 57398 | 12 | imports BNF_Cardinal_Arithmetic Fun_Def_Base | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 13 | keywords | 
| 49286 | 14 | "print_bnfs" :: diag and | 
| 69913 | 15 | "bnf" :: thy_goal_defn | 
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 16 | begin | 
| 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 17 | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61423diff
changeset | 18 | lemma Collect_case_prodD: "x \<in> Collect (case_prod A) \<Longrightarrow> A (fst x) (snd x)" | 
| 58104 | 19 | by auto | 
| 20 | ||
| 58916 | 21 | inductive | 
| 22 |    rel_sum :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd \<Rightarrow> bool" for R1 R2
 | |
| 58446 | 23 | where | 
| 58916 | 24 | "R1 a c \<Longrightarrow> rel_sum R1 R2 (Inl a) (Inl c)" | 
| 25 | | "R2 b d \<Longrightarrow> rel_sum R1 R2 (Inr b) (Inr d)" | |
| 26 | ||
| 58446 | 27 | definition | 
| 57398 | 28 |   rel_fun :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> bool"
 | 
| 29 | where | |
| 30 | "rel_fun A B = (\<lambda>f g. \<forall>x y. A x y \<longrightarrow> B (f x) (g y))" | |
| 31 | ||
| 32 | lemma rel_funI [intro]: | |
| 33 | assumes "\<And>x y. A x y \<Longrightarrow> B (f x) (g y)" | |
| 34 | shows "rel_fun A B f g" | |
| 35 | using assms by (simp add: rel_fun_def) | |
| 36 | ||
| 37 | lemma rel_funD: | |
| 38 | assumes "rel_fun A B f g" and "A x y" | |
| 39 | shows "B (f x) (g y)" | |
| 40 | using assms by (simp add: rel_fun_def) | |
| 41 | ||
| 59513 | 42 | lemma rel_fun_mono: | 
| 43 | "\<lbrakk> rel_fun X A f g; \<And>x y. Y x y \<longrightarrow> X x y; \<And>x y. A x y \<Longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_fun Y B f g" | |
| 44 | by(simp add: rel_fun_def) | |
| 45 | ||
| 46 | lemma rel_fun_mono' [mono]: | |
| 47 | "\<lbrakk> \<And>x y. Y x y \<longrightarrow> X x y; \<And>x y. A x y \<longrightarrow> B x y \<rbrakk> \<Longrightarrow> rel_fun X A f g \<longrightarrow> rel_fun Y B f g" | |
| 48 | by(simp add: rel_fun_def) | |
| 49 | ||
| 58104 | 50 | definition rel_set :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool"
 | 
| 51 | where "rel_set R = (\<lambda>A B. (\<forall>x\<in>A. \<exists>y\<in>B. R x y) \<and> (\<forall>y\<in>B. \<exists>x\<in>A. R x y))" | |
| 52 | ||
| 53 | lemma rel_setI: | |
| 54 | assumes "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. R x y" | |
| 55 | assumes "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. R x y" | |
| 56 | shows "rel_set R A B" | |
| 57 | using assms unfolding rel_set_def by simp | |
| 58 | ||
| 59 | lemma predicate2_transferD: | |
| 67399 | 60 |    "\<lbrakk>rel_fun R1 (rel_fun R2 (=)) P Q; a \<in> A; b \<in> B; A \<subseteq> {(x, y). R1 x y}; B \<subseteq> {(x, y). R2 x y}\<rbrakk> \<Longrightarrow>
 | 
| 58104 | 61 | P (fst a) (fst b) \<longleftrightarrow> Q (snd a) (snd b)" | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61423diff
changeset | 62 | unfolding rel_fun_def by (blast dest!: Collect_case_prodD) | 
| 58104 | 63 | |
| 57398 | 64 | definition collect where | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 65 | "collect F x = (\<Union>f \<in> F. f x)" | 
| 57398 | 66 | |
| 67 | lemma fstI: "x = (y, z) \<Longrightarrow> fst x = y" | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 68 | by simp | 
| 57398 | 69 | |
| 70 | lemma sndI: "x = (y, z) \<Longrightarrow> snd x = z" | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 71 | by simp | 
| 57398 | 72 | |
| 73 | lemma bijI': "\<lbrakk>\<And>x y. (f x = f y) = (x = y); \<And>y. \<exists>x. y = f x\<rbrakk> \<Longrightarrow> bij f" | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 74 | unfolding bij_def inj_on_def by auto blast | 
| 57398 | 75 | |
| 76 | (* Operator: *) | |
| 77 | definition "Gr A f = {(a, f a) | a. a \<in> A}"
 | |
| 78 | ||
| 79 | definition "Grp A f = (\<lambda>a b. b = f a \<and> a \<in> A)" | |
| 80 | ||
| 81 | definition vimage2p where | |
| 82 | "vimage2p f g R = (\<lambda>x y. R (f x) (g y))" | |
| 83 | ||
| 56635 | 84 | lemma collect_comp: "collect F \<circ> g = collect ((\<lambda>f. f \<circ> g) ` F)" | 
| 66198 | 85 | by (rule ext) (simp add: collect_def) | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 86 | |
| 57641 
dc59f147b27d
more robust notation BNF_Def.convol, which is private to main HOL, but may cause syntax ambiguities nonetheless (e.g. List.thy);
 wenzelm parents: 
57398diff
changeset | 87 | definition convol ("\<langle>(_,/ _)\<rangle>") where
 | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 88 | "\<langle>f, g\<rangle> \<equiv> \<lambda>a. (f a, g a)" | 
| 49495 | 89 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 90 | lemma fst_convol: "fst \<circ> \<langle>f, g\<rangle> = f" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 91 | apply(rule ext) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 92 | unfolding convol_def by simp | 
| 49495 | 93 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 94 | lemma snd_convol: "snd \<circ> \<langle>f, g\<rangle> = g" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 95 | apply(rule ext) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 96 | unfolding convol_def by simp | 
| 49495 | 97 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 98 | lemma convol_mem_GrpI: | 
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60758diff
changeset | 99 | "x \<in> A \<Longrightarrow> \<langle>id, g\<rangle> x \<in> (Collect (case_prod (Grp A g)))" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 100 | unfolding convol_def Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 101 | |
| 49312 | 102 | definition csquare where | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 103 | "csquare A f1 f2 p1 p2 \<longleftrightarrow> (\<forall> a \<in> A. f1 (p1 a) = f2 (p2 a))" | 
| 49312 | 104 | |
| 67399 | 105 | lemma eq_alt: "(=) = Grp UNIV id" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 106 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 107 | |
| 67613 | 108 | lemma leq_conversepI: "R = (=) \<Longrightarrow> R \<le> R\<inverse>\<inverse>" | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 109 | by auto | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 110 | |
| 67399 | 111 | lemma leq_OOI: "R = (=) \<Longrightarrow> R \<le> R OO R" | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 112 | by auto | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 113 | |
| 67613 | 114 | lemma OO_Grp_alt: "(Grp A f)\<inverse>\<inverse> OO Grp A g = (\<lambda>x y. \<exists>z. z \<in> A \<and> f z = x \<and> g z = y)" | 
| 53561 | 115 | unfolding Grp_def by auto | 
| 116 | ||
| 67613 | 117 | lemma Grp_UNIV_id: "f = id \<Longrightarrow> (Grp UNIV f)\<inverse>\<inverse> OO Grp UNIV f = Grp UNIV f" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 118 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 119 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 120 | lemma Grp_UNIV_idI: "x = y \<Longrightarrow> Grp UNIV id x y" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 121 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 122 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 123 | lemma Grp_mono: "A \<le> B \<Longrightarrow> Grp A f \<le> Grp B f" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 124 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 125 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 126 | lemma GrpI: "\<lbrakk>f x = y; x \<in> A\<rbrakk> \<Longrightarrow> Grp A f x y" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 127 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 128 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 129 | lemma GrpE: "Grp A f x y \<Longrightarrow> (\<lbrakk>f x = y; x \<in> A\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 130 | unfolding Grp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 131 | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61423diff
changeset | 132 | lemma Collect_case_prod_Grp_eqD: "z \<in> Collect (case_prod (Grp A f)) \<Longrightarrow> (f \<circ> fst) z = snd z" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 133 | unfolding Grp_def comp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 134 | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61423diff
changeset | 135 | lemma Collect_case_prod_Grp_in: "z \<in> Collect (case_prod (Grp A f)) \<Longrightarrow> fst z \<in> A" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 136 | unfolding Grp_def comp_def by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 137 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 138 | definition "pick_middlep P Q a c = (SOME b. P a b \<and> Q b c)" | 
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 139 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 140 | lemma pick_middlep: | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 141 | "(P OO Q) a c \<Longrightarrow> P a (pick_middlep P Q a c) \<and> Q (pick_middlep P Q a c) c" | 
| 74886 | 142 | unfolding pick_middlep_def by (rule someI_ex) auto | 
| 49312 | 143 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 144 | definition fstOp where | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 145 | "fstOp P Q ac = (fst ac, pick_middlep P Q (fst ac) (snd ac))" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 146 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 147 | definition sndOp where | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 148 | "sndOp P Q ac = (pick_middlep P Q (fst ac) (snd ac), (snd ac))" | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 149 | |
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60758diff
changeset | 150 | lemma fstOp_in: "ac \<in> Collect (case_prod (P OO Q)) \<Longrightarrow> fstOp P Q ac \<in> Collect (case_prod P)" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 151 | unfolding fstOp_def mem_Collect_eq | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 152 | by (subst (asm) surjective_pairing, unfold prod.case) (erule pick_middlep[THEN conjunct1]) | 
| 49312 | 153 | |
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 154 | lemma fst_fstOp: "fst bc = (fst \<circ> fstOp P Q) bc" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 155 | unfolding comp_def fstOp_def by simp | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 156 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 157 | lemma snd_sndOp: "snd bc = (snd \<circ> sndOp P Q) bc" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 158 | unfolding comp_def sndOp_def by simp | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 159 | |
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60758diff
changeset | 160 | lemma sndOp_in: "ac \<in> Collect (case_prod (P OO Q)) \<Longrightarrow> sndOp P Q ac \<in> Collect (case_prod Q)" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 161 | unfolding sndOp_def mem_Collect_eq | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 162 | by (subst (asm) surjective_pairing, unfold prod.case) (erule pick_middlep[THEN conjunct2]) | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 163 | |
| 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 164 | lemma csquare_fstOp_sndOp: | 
| 61423 | 165 | "csquare (Collect (f (P OO Q))) snd fst (fstOp P Q) (sndOp P Q)" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 166 | unfolding csquare_def fstOp_def sndOp_def using pick_middlep by simp | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 167 | |
| 56635 | 168 | lemma snd_fst_flip: "snd xy = (fst \<circ> (%(x, y). (y, x))) xy" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 169 | by (simp split: prod.split) | 
| 49312 | 170 | |
| 56635 | 171 | lemma fst_snd_flip: "fst xy = (snd \<circ> (%(x, y). (y, x))) xy" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 172 | by (simp split: prod.split) | 
| 49312 | 173 | |
| 67613 | 174 | lemma flip_pred: "A \<subseteq> Collect (case_prod (R \<inverse>\<inverse>)) \<Longrightarrow> (%(x, y). (y, x)) ` A \<subseteq> Collect (case_prod R)" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 175 | by auto | 
| 51893 
596baae88a88
got rid of the set based relator---use (binary) predicate based relator instead
 traytel parents: 
51836diff
changeset | 176 | |
| 51917 
f964a9887713
store proper theorems even for fixed points that have no passive live variables
 traytel parents: 
51916diff
changeset | 177 | lemma predicate2_eqD: "A = B \<Longrightarrow> A a b \<longleftrightarrow> B a b" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 178 | by simp | 
| 49537 
fe1deee434b6
generate "rel_as_srel" and "rel_flip" properties
 blanchet parents: 
49510diff
changeset | 179 | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 180 | lemma case_sum_o_inj: "case_sum f g \<circ> Inl = f" "case_sum f g \<circ> Inr = g" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 181 | by auto | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 182 | |
| 67091 | 183 | lemma map_sum_o_inj: "map_sum f g \<circ> Inl = Inl \<circ> f" "map_sum f g \<circ> Inr = Inr \<circ> g" | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 184 | by auto | 
| 57802 | 185 | |
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 186 | lemma card_order_csum_cone_cexp_def: | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 187 |   "card_order r \<Longrightarrow> ( |A1| +c cone) ^c r = |Func UNIV (Inl ` A1 \<union> {Inr ()})|"
 | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 188 | unfolding cexp_def cone_def Field_csum Field_card_of by (auto dest: Field_card_order) | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 189 | |
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 190 | lemma If_the_inv_into_in_Func: | 
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 191 |   "\<lbrakk>inj_on g C; C \<subseteq> B \<union> {x}\<rbrakk> \<Longrightarrow>
 | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 192 |    (\<lambda>i. if i \<in> g ` C then the_inv_into C g i else x) \<in> Func UNIV (B \<union> {x})"
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 193 | unfolding Func_def by (auto dest: the_inv_into_into) | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 194 | |
| 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 195 | lemma If_the_inv_into_f_f: | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 196 | "\<lbrakk>i \<in> C; inj_on g C\<rbrakk> \<Longrightarrow> ((\<lambda>i. if i \<in> g ` C then the_inv_into C g i else x) \<circ> g) i = id i" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 197 | unfolding Func_def by (auto elim: the_inv_into_f_f) | 
| 52635 
4f84b730c489
got rid of in_bd BNF property (derivable from set_bd+map_cong+map_comp+map_id)
 traytel parents: 
51917diff
changeset | 198 | |
| 56635 | 199 | lemma the_inv_f_o_f_id: "inj f \<Longrightarrow> (the_inv f \<circ> f) z = id z" | 
| 200 | by (simp add: the_inv_f_f) | |
| 201 | ||
| 52731 | 202 | lemma vimage2pI: "R (f x) (g y) \<Longrightarrow> vimage2p f g R x y" | 
| 63834 | 203 | unfolding vimage2p_def . | 
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52660diff
changeset | 204 | |
| 55945 | 205 | lemma rel_fun_iff_leq_vimage2p: "(rel_fun R S) f g = (R \<le> vimage2p f g S)" | 
| 206 | unfolding rel_fun_def vimage2p_def by auto | |
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52660diff
changeset | 207 | |
| 61032 
b57df8eecad6
standardized some occurences of ancient "split" alias
 haftmann parents: 
60758diff
changeset | 208 | lemma convol_image_vimage2p: "\<langle>f \<circ> fst, g \<circ> snd\<rangle> ` Collect (case_prod (vimage2p f g R)) \<subseteq> Collect (case_prod R)" | 
| 52731 | 209 | unfolding vimage2p_def convol_def by auto | 
| 52719 
480a3479fa47
transfer rule for map (not yet registered as a transfer rule)
 traytel parents: 
52660diff
changeset | 210 | |
| 54961 | 211 | lemma vimage2p_Grp: "vimage2p f g P = Grp UNIV f OO P OO (Grp UNIV g)\<inverse>\<inverse>" | 
| 212 | unfolding vimage2p_def Grp_def by auto | |
| 213 | ||
| 58106 | 214 | lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)" | 
| 215 | by simp | |
| 216 | ||
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 217 | lemma comp_apply_eq: "f (g x) = h (k x) \<Longrightarrow> (f \<circ> g) x = (h \<circ> k) x" | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 218 | unfolding comp_apply by assumption | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: 
58106diff
changeset | 219 | |
| 67399 | 220 | lemma refl_ge_eq: "(\<And>x. R x x) \<Longrightarrow> (=) \<le> R" | 
| 59726 | 221 | by auto | 
| 222 | ||
| 67399 | 223 | lemma ge_eq_refl: "(=) \<le> R \<Longrightarrow> R x x" | 
| 59726 | 224 | by auto | 
| 225 | ||
| 67399 | 226 | lemma reflp_eq: "reflp R = ((=) \<le> R)" | 
| 61240 | 227 | by (auto simp: reflp_def fun_eq_iff) | 
| 228 | ||
| 229 | lemma transp_relcompp: "transp r \<longleftrightarrow> r OO r \<le> r" | |
| 230 | by (auto simp: transp_def) | |
| 231 | ||
| 232 | lemma symp_conversep: "symp R = (R\<inverse>\<inverse> \<le> R)" | |
| 233 | by (auto simp: symp_def fun_eq_iff) | |
| 234 | ||
| 235 | lemma diag_imp_eq_le: "(\<And>x. x \<in> A \<Longrightarrow> R x x) \<Longrightarrow> \<forall>x y. x \<in> A \<longrightarrow> y \<in> A \<longrightarrow> x = y \<longrightarrow> R x y" | |
| 236 | by blast | |
| 237 | ||
| 62324 | 238 | definition eq_onp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 239 | where "eq_onp R = (\<lambda>x y. R x \<and> x = y)" | |
| 240 | ||
| 241 | lemma eq_onp_Grp: "eq_onp P = BNF_Def.Grp (Collect P) id" | |
| 242 | unfolding eq_onp_def Grp_def by auto | |
| 243 | ||
| 244 | lemma eq_onp_to_eq: "eq_onp P x y \<Longrightarrow> x = y" | |
| 245 | by (simp add: eq_onp_def) | |
| 246 | ||
| 67399 | 247 | lemma eq_onp_top_eq_eq: "eq_onp top = (=)" | 
| 62324 | 248 | by (simp add: eq_onp_def) | 
| 249 | ||
| 250 | lemma eq_onp_same_args: "eq_onp P x x = P x" | |
| 63092 | 251 | by (auto simp add: eq_onp_def) | 
| 62324 | 252 | |
| 253 | lemma eq_onp_eqD: "eq_onp P = Q \<Longrightarrow> P x = Q x x" | |
| 254 | unfolding eq_onp_def by blast | |
| 255 | ||
| 256 | lemma Ball_Collect: "Ball A P = (A \<subseteq> (Collect P))" | |
| 257 | by auto | |
| 258 | ||
| 259 | lemma eq_onp_mono0: "\<forall>x\<in>A. P x \<longrightarrow> Q x \<Longrightarrow> \<forall>x\<in>A. \<forall>y\<in>A. eq_onp P x y \<longrightarrow> eq_onp Q x y" | |
| 260 | unfolding eq_onp_def by auto | |
| 261 | ||
| 67399 | 262 | lemma eq_onp_True: "eq_onp (\<lambda>_. True) = (=)" | 
| 62324 | 263 | unfolding eq_onp_def by simp | 
| 264 | ||
| 67091 | 265 | lemma Ball_image_comp: "Ball (f ` A) g = Ball A (g \<circ> f)" | 
| 62324 | 266 | by auto | 
| 267 | ||
| 62329 | 268 | lemma rel_fun_Collect_case_prodD: | 
| 67091 | 269 | "rel_fun A B f g \<Longrightarrow> X \<subseteq> Collect (case_prod A) \<Longrightarrow> x \<in> X \<Longrightarrow> B ((f \<circ> fst) x) ((g \<circ> snd) x)" | 
| 62329 | 270 | unfolding rel_fun_def by auto | 
| 271 | ||
| 63714 | 272 | lemma eq_onp_mono_iff: "eq_onp P \<le> eq_onp Q \<longleftrightarrow> P \<le> Q" | 
| 273 | unfolding eq_onp_def by auto | |
| 274 | ||
| 69605 | 275 | ML_file \<open>Tools/BNF/bnf_util.ML\<close> | 
| 276 | ML_file \<open>Tools/BNF/bnf_tactics.ML\<close> | |
| 277 | ML_file \<open>Tools/BNF/bnf_def_tactics.ML\<close> | |
| 278 | ML_file \<open>Tools/BNF/bnf_def.ML\<close> | |
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: 
49286diff
changeset | 279 | |
| 48975 
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
 blanchet parents: diff
changeset | 280 | end |