author | blanchet |
Wed, 05 Feb 2014 11:22:36 +0100 | |
changeset 55332 | 803a7400cc58 |
parent 55083 | 0a689157e3ce |
child 55414 | eab03e9cee8a |
permissions | -rw-r--r-- |
55059 | 1 |
(* Title: HOL/BNF_FP_Base.thy |
53311 | 2 |
Author: Lorenz Panny, TU Muenchen |
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
3 |
Author: Dmitriy Traytel, TU Muenchen |
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
4 |
Author: Jasmin Blanchette, TU Muenchen |
53311 | 5 |
Copyright 2012, 2013 |
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
6 |
|
55059 | 7 |
Shared fixed point operations on bounded natural functors. |
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
8 |
*) |
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
9 |
|
53311 | 10 |
header {* Shared Fixed Point Operations on Bounded Natural Functors *} |
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
11 |
|
53311 | 12 |
theory BNF_FP_Base |
55083 | 13 |
imports BNF_Comp |
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
14 |
begin |
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
15 |
|
49590 | 16 |
lemma mp_conj: "(P \<longrightarrow> Q) \<and> R \<Longrightarrow> P \<Longrightarrow> R \<and> Q" |
17 |
by auto |
|
18 |
||
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49590
diff
changeset
|
19 |
lemma eq_sym_Unity_conv: "(x = (() = ())) = x" |
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
20 |
by blast |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
21 |
|
53903 | 22 |
lemma unit_case_Unity: "(case u of () \<Rightarrow> f) = f" |
49312 | 23 |
by (cases u) (hypsubst, rule unit.cases) |
24 |
||
49539
be6cbf960aa7
fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents:
49510
diff
changeset
|
25 |
lemma prod_case_Pair_iden: "(case p of (x, y) \<Rightarrow> (x, y)) = p" |
be6cbf960aa7
fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents:
49510
diff
changeset
|
26 |
by simp |
be6cbf960aa7
fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents:
49510
diff
changeset
|
27 |
|
49335 | 28 |
lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x" |
29 |
by simp |
|
30 |
||
31 |
lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x" |
|
32 |
by clarify |
|
33 |
||
34 |
lemma prod_all_impI_step: "(\<And>x. \<forall>y. P (x, y) \<longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x" |
|
35 |
by auto |
|
36 |
||
49683 | 37 |
lemma pointfree_idE: "f \<circ> g = id \<Longrightarrow> f (g x) = x" |
55066 | 38 |
unfolding comp_def fun_eq_iff by simp |
49312 | 39 |
|
40 |
lemma o_bij: |
|
49683 | 41 |
assumes gf: "g \<circ> f = id" and fg: "f \<circ> g = id" |
49312 | 42 |
shows "bij f" |
43 |
unfolding bij_def inj_on_def surj_def proof safe |
|
44 |
fix a1 a2 assume "f a1 = f a2" |
|
45 |
hence "g ( f a1) = g (f a2)" by simp |
|
46 |
thus "a1 = a2" using gf unfolding fun_eq_iff by simp |
|
47 |
next |
|
48 |
fix b |
|
49 |
have "b = f (g b)" |
|
50 |
using fg unfolding fun_eq_iff by simp |
|
51 |
thus "EX a. b = f a" by blast |
|
52 |
qed |
|
53 |
||
54 |
lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp |
|
55 |
||
56 |
lemma sum_case_step: |
|
49683 | 57 |
"sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p" |
58 |
"sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p" |
|
49312 | 59 |
by auto |
60 |
||
61 |
lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
|
62 |
by simp |
|
63 |
||
64 |
lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P" |
|
65 |
by blast |
|
66 |
||
67 |
lemma obj_sumE_f: |
|
68 |
"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f x \<longrightarrow> P" |
|
52660 | 69 |
by (rule allI) (metis sumE) |
49312 | 70 |
|
71 |
lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P" |
|
72 |
by (cases s) auto |
|
73 |
||
74 |
lemma sum_case_if: |
|
75 |
"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)" |
|
76 |
by simp |
|
77 |
||
49428
93f158d59ead
handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents:
49427
diff
changeset
|
78 |
lemma mem_UN_compreh_eq: "(z : \<Union>{y. \<exists>x\<in>A. y = F x}) = (\<exists>x\<in>A. z : F x)" |
93f158d59ead
handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents:
49427
diff
changeset
|
79 |
by blast |
93f158d59ead
handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents:
49427
diff
changeset
|
80 |
|
49585
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
81 |
lemma UN_compreh_eq_eq: |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
82 |
"\<Union>{y. \<exists>x\<in>A. y = {}} = {}" |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
83 |
"\<Union>{y. \<exists>x\<in>A. y = {x}} = A" |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
84 |
by blast+ |
5c4a12550491
generate high-level "maps", "sets", and "rels" properties
blanchet
parents:
49539
diff
changeset
|
85 |
|
51745
a06a3c777add
simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents:
51740
diff
changeset
|
86 |
lemma Inl_Inr_False: "(Inl x = Inr y) = False" |
a06a3c777add
simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents:
51740
diff
changeset
|
87 |
by simp |
a06a3c777add
simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents:
51740
diff
changeset
|
88 |
|
49429
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
89 |
lemma prod_set_simps: |
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
90 |
"fsts (x, y) = {x}" |
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
91 |
"snds (x, y) = {y}" |
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
92 |
unfolding fsts_def snds_def by simp+ |
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
93 |
|
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
94 |
lemma sum_set_simps: |
49451
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49430
diff
changeset
|
95 |
"setl (Inl x) = {x}" |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49430
diff
changeset
|
96 |
"setl (Inr x) = {}" |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49430
diff
changeset
|
97 |
"setr (Inl x) = {}" |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49430
diff
changeset
|
98 |
"setr (Inr x) = {x}" |
7a28d22c33c6
renamed "sum_setl" to "setl" and similarly for r
blanchet
parents:
49430
diff
changeset
|
99 |
unfolding sum_set_defs by simp+ |
49429
64ac3471005a
cleaner way of dealing with the set functions of sums and products
blanchet
parents:
49428
diff
changeset
|
100 |
|
52505
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents:
52349
diff
changeset
|
101 |
lemma spec2: "\<forall>x y. P x y \<Longrightarrow> P x y" |
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents:
52349
diff
changeset
|
102 |
by blast |
e62f3fd2035e
share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents:
52349
diff
changeset
|
103 |
|
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
104 |
lemma rewriteR_comp_comp: "\<lbrakk>g o h = r\<rbrakk> \<Longrightarrow> f o g o h = f o r" |
55066 | 105 |
unfolding comp_def fun_eq_iff by auto |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
106 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
107 |
lemma rewriteR_comp_comp2: "\<lbrakk>g o h = r1 o r2; f o r1 = l\<rbrakk> \<Longrightarrow> f o g o h = l o r2" |
55066 | 108 |
unfolding comp_def fun_eq_iff by auto |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
109 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
110 |
lemma rewriteL_comp_comp: "\<lbrakk>f o g = l\<rbrakk> \<Longrightarrow> f o (g o h) = l o h" |
55066 | 111 |
unfolding comp_def fun_eq_iff by auto |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
112 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
113 |
lemma rewriteL_comp_comp2: "\<lbrakk>f o g = l1 o l2; l2 o h = r\<rbrakk> \<Longrightarrow> f o (g o h) = l1 o r" |
55066 | 114 |
unfolding comp_def fun_eq_iff by auto |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
115 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
116 |
lemma convol_o: "<f, g> o h = <f o h, g o h>" |
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
117 |
unfolding convol_def by auto |
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
118 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
119 |
lemma map_pair_o_convol: "map_pair h1 h2 o <f, g> = <h1 o f, h2 o g>" |
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
120 |
unfolding convol_def by auto |
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
121 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
122 |
lemma map_pair_o_convol_id: "(map_pair f id \<circ> <id , g>) x = <id \<circ> f , g> x" |
55066 | 123 |
unfolding map_pair_o_convol id_comp comp_id .. |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
124 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
125 |
lemma o_sum_case: "h o sum_case f g = sum_case (h o f) (h o g)" |
55066 | 126 |
unfolding comp_def by (auto split: sum.splits) |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
127 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
128 |
lemma sum_case_o_sum_map: "sum_case f g o sum_map h1 h2 = sum_case (f o h1) (g o h2)" |
55066 | 129 |
unfolding comp_def by (auto split: sum.splits) |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
130 |
|
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
131 |
lemma sum_case_o_sum_map_id: "(sum_case id g o sum_map f id) x = sum_case (f o id) g x" |
55066 | 132 |
unfolding sum_case_o_sum_map id_comp comp_id .. |
52913
2d2d9d1de1a9
theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents:
52731
diff
changeset
|
133 |
|
52731 | 134 |
lemma fun_rel_def_butlast: |
135 |
"(fun_rel R (fun_rel S T)) f g = (\<forall>x y. R x y \<longrightarrow> (fun_rel S T) (f x) (g y))" |
|
136 |
unfolding fun_rel_def .. |
|
137 |
||
53105
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
138 |
lemma subst_eq_imp: "(\<forall>a b. a = b \<longrightarrow> P a b) \<equiv> (\<forall>a. P a a)" |
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
139 |
by auto |
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
140 |
|
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
141 |
lemma eq_subset: "op = \<le> (\<lambda>a b. P a b \<or> a = b)" |
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
142 |
by auto |
ec38e9f4352f
simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents:
52913
diff
changeset
|
143 |
|
53308 | 144 |
lemma eq_le_Grp_id_iff: "(op = \<le> Grp (Collect R) id) = (All R)" |
145 |
unfolding Grp_def id_apply by blast |
|
146 |
||
147 |
lemma Grp_id_mono_subst: "(\<And>x y. Grp P id x y \<Longrightarrow> Grp Q id (f x) (f y)) \<equiv> |
|
148 |
(\<And>x. x \<in> P \<Longrightarrow> f x \<in> Q)" |
|
149 |
unfolding Grp_def by rule auto |
|
150 |
||
55062 | 151 |
ML_file "Tools/BNF/bnf_fp_util.ML" |
152 |
ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML" |
|
153 |
ML_file "Tools/BNF/bnf_fp_def_sugar.ML" |
|
154 |
ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML" |
|
155 |
ML_file "Tools/BNF/bnf_fp_n2m.ML" |
|
156 |
ML_file "Tools/BNF/bnf_fp_n2m_sugar.ML" |
|
157 |
ML_file "Tools/BNF/bnf_fp_rec_sugar_util.ML" |
|
49309
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents:
49308
diff
changeset
|
158 |
|
49308
6190b701e4f4
reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff
changeset
|
159 |
end |