src/HOL/BNF_Fixpoint_Base.thy
author wenzelm
Mon, 20 Jun 2016 21:40:48 +0200
changeset 63325 1086d56cde86
parent 63046 8053ef5a0174
child 67091 1393c2340eec
permissions -rw-r--r--
misc tuning and modernization;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58128
43a1ba26a8cb renamed BNF theories
blanchet
parents: 58112
diff changeset
     1
(*  Title:      HOL/BNF_Fixpoint_Base.thy
53311
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
     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
57698
afef6616cbae header tuning
blanchet
parents: 57641
diff changeset
     5
    Author:     Martin Desharnais, TU Muenchen
afef6616cbae header tuning
blanchet
parents: 57641
diff changeset
     6
    Copyright   2012, 2013, 2014
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
     7
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents: 58179
diff changeset
     8
Shared fixpoint 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
     9
*)
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    10
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 58916
diff changeset
    11
section \<open>Shared Fixpoint Operations on Bounded Natural Functors\<close>
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    12
58128
43a1ba26a8cb renamed BNF theories
blanchet
parents: 58112
diff changeset
    13
theory BNF_Fixpoint_Base
43a1ba26a8cb renamed BNF theories
blanchet
parents: 58112
diff changeset
    14
imports BNF_Composition Basic_BNFs
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    15
begin
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    16
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57489
diff changeset
    17
lemma conj_imp_eq_imp_imp: "(P \<and> Q \<Longrightarrow> PROP R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> PROP R)"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60918
diff changeset
    18
  by standard simp_all
57525
f9dd8a33f820 generate 'rel_cases' theorem for (co)datatypes
desharna
parents: 57489
diff changeset
    19
62158
c25c62055180 generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
blanchet
parents: 61551
diff changeset
    20
lemma predicate2D_conj: "P \<le> Q \<and> R \<Longrightarrow> R \<and> (P x y \<longrightarrow> Q x y)"
c25c62055180 generate stronger 'rel_(co)induct' and 'coinduct' principles for mutually (co)recursive (co)datatypes
blanchet
parents: 61551
diff changeset
    21
  by blast
57302
58f02fbaa764 generate 'rel_coinduct' theorem for codatatypes
desharna
parents: 57301
diff changeset
    22
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    23
lemma eq_sym_Unity_conv: "(x = (() = ())) = x"
58159
blanchet
parents: 58128
diff changeset
    24
  by blast
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    25
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    26
lemma case_unit_Unity: "(case u of () \<Rightarrow> f) = f"
58159
blanchet
parents: 58128
diff changeset
    27
  by (cases u) (hypsubst, rule unit.case)
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    28
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    29
lemma case_prod_Pair_iden: "(case p of (x, y) \<Rightarrow> (x, y)) = p"
58159
blanchet
parents: 58128
diff changeset
    30
  by simp
49539
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    31
49335
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    32
lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
58159
blanchet
parents: 58128
diff changeset
    33
  by simp
49335
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    34
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    35
lemma pointfree_idE: "f \<circ> g = id \<Longrightarrow> f (g x) = x"
58159
blanchet
parents: 58128
diff changeset
    36
  unfolding comp_def fun_eq_iff by simp
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    37
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    38
lemma o_bij:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    39
  assumes gf: "g \<circ> f = id" and fg: "f \<circ> g = id"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    40
  shows "bij f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    41
unfolding bij_def inj_on_def surj_def proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    42
  fix a1 a2 assume "f a1 = f a2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
  hence "g ( f a1) = g (f a2)" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    44
  thus "a1 = a2" using gf unfolding fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    45
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
  fix b
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
  have "b = f (g b)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
  using fg unfolding fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    49
  thus "EX a. b = f a" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    50
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    51
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    52
lemma case_sum_step:
58159
blanchet
parents: 58128
diff changeset
    53
  "case_sum (case_sum f' g') g (Inl p) = case_sum f' g' p"
blanchet
parents: 58128
diff changeset
    54
  "case_sum f (case_sum f' g') (Inr p) = case_sum f' g' p"
blanchet
parents: 58128
diff changeset
    55
  by auto
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    56
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    57
lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
58159
blanchet
parents: 58128
diff changeset
    58
  by blast
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    59
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
    60
lemma type_copy_obj_one_point_absE:
55811
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    61
  assumes "type_definition Rep Abs UNIV" "\<forall>x. s = Abs x \<longrightarrow> P" shows P
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    62
  using type_definition.Rep_inverse[OF assms(1)]
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    63
  by (intro mp[OF spec[OF assms(2), of "Rep s"]]) simp
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
    64
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    65
lemma obj_sumE_f:
55811
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    66
  assumes "\<forall>x. s = f (Inl x) \<longrightarrow> P" "\<forall>x. s = f (Inr x) \<longrightarrow> P"
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    67
  shows "\<forall>x. s = f x \<longrightarrow> P"
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    68
proof
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    69
  fix x from assms show "s = f x \<longrightarrow> P" by (cases x) auto
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    70
qed
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    71
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    72
lemma case_sum_if:
58159
blanchet
parents: 58128
diff changeset
    73
  "case_sum f g (if p then Inl x else Inr y) = (if p then f x else g y)"
blanchet
parents: 58128
diff changeset
    74
  by simp
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    75
62325
7e4d31eefe60 simp rules for fsts, snds, setl, setr
traytel
parents: 62158
diff changeset
    76
lemma prod_set_simps[simp]:
58159
blanchet
parents: 58128
diff changeset
    77
  "fsts (x, y) = {x}"
blanchet
parents: 58128
diff changeset
    78
  "snds (x, y) = {y}"
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58889
diff changeset
    79
  unfolding prod_set_defs by simp+
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    80
62325
7e4d31eefe60 simp rules for fsts, snds, setl, setr
traytel
parents: 62158
diff changeset
    81
lemma sum_set_simps[simp]:
58159
blanchet
parents: 58128
diff changeset
    82
  "setl (Inl x) = {x}"
blanchet
parents: 58128
diff changeset
    83
  "setl (Inr x) = {}"
blanchet
parents: 58128
diff changeset
    84
  "setr (Inl x) = {}"
blanchet
parents: 58128
diff changeset
    85
  "setr (Inr x) = {x}"
blanchet
parents: 58128
diff changeset
    86
  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
    87
57301
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57279
diff changeset
    88
lemma Inl_Inr_False: "(Inl x = Inr y) = False"
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57279
diff changeset
    89
  by simp
7b997028aaac generate 'rel_coinduct0' theorem for codatatypes
desharna
parents: 57279
diff changeset
    90
57471
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57303
diff changeset
    91
lemma Inr_Inl_False: "(Inr x = Inl y) = False"
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57303
diff changeset
    92
  by simp
11cd462e31ec generate 'rel_induct' theorem for datatypes
desharna
parents: 57303
diff changeset
    93
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
    94
lemma spec2: "\<forall>x y. P x y \<Longrightarrow> P x y"
58159
blanchet
parents: 58128
diff changeset
    95
  by blast
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
    96
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
    97
lemma rewriteR_comp_comp: "\<lbrakk>g \<circ> h = r\<rbrakk> \<Longrightarrow> f \<circ> g \<circ> h = f \<circ> r"
55066
blanchet
parents: 55062
diff changeset
    98
  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
    99
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   100
lemma rewriteR_comp_comp2: "\<lbrakk>g \<circ> h = r1 \<circ> r2; f \<circ> r1 = l\<rbrakk> \<Longrightarrow> f \<circ> g \<circ> h = l \<circ> r2"
55066
blanchet
parents: 55062
diff changeset
   101
  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
   102
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   103
lemma rewriteL_comp_comp: "\<lbrakk>f \<circ> g = l\<rbrakk> \<Longrightarrow> f \<circ> (g \<circ> h) = l \<circ> h"
55066
blanchet
parents: 55062
diff changeset
   104
  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
   105
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   106
lemma rewriteL_comp_comp2: "\<lbrakk>f \<circ> g = l1 \<circ> l2; l2 \<circ> h = r\<rbrakk> \<Longrightarrow> f \<circ> (g \<circ> h) = l1 \<circ> r"
55066
blanchet
parents: 55062
diff changeset
   107
  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
   108
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: 57525
diff changeset
   109
lemma convol_o: "\<langle>f, g\<rangle> \<circ> h = \<langle>f \<circ> h, g \<circ> h\<rangle>"
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   110
  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
   111
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: 57525
diff changeset
   112
lemma map_prod_o_convol: "map_prod h1 h2 \<circ> \<langle>f, g\<rangle> = \<langle>h1 \<circ> f, h2 \<circ> g\<rangle>"
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   113
  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
   114
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: 57525
diff changeset
   115
lemma map_prod_o_convol_id: "(map_prod f id \<circ> \<langle>id, g\<rangle>) x = \<langle>id \<circ> f, g\<rangle> x"
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55931
diff changeset
   116
  unfolding map_prod_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
   117
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   118
lemma o_case_sum: "h \<circ> case_sum f g = case_sum (h \<circ> f) (h \<circ> g)"
55066
blanchet
parents: 55062
diff changeset
   119
  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
   120
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   121
lemma case_sum_o_map_sum: "case_sum f g \<circ> map_sum h1 h2 = case_sum (f \<circ> h1) (g \<circ> h2)"
55066
blanchet
parents: 55062
diff changeset
   122
  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
   123
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   124
lemma case_sum_o_map_sum_id: "(case_sum id g \<circ> map_sum f id) x = case_sum (f \<circ> id) g x"
55931
62156e694f3d renamed 'map_sum' to 'sum_map'
blanchet
parents: 55930
diff changeset
   125
  unfolding case_sum_o_map_sum 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
   126
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55936
diff changeset
   127
lemma rel_fun_def_butlast:
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55936
diff changeset
   128
  "rel_fun R (rel_fun S T) f g = (\<forall>x y. R x y \<longrightarrow> (rel_fun S T) (f x) (g y))"
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55936
diff changeset
   129
  unfolding rel_fun_def ..
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   130
53105
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   131
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
   132
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   133
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   134
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
   135
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   136
53308
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   137
lemma eq_le_Grp_id_iff: "(op = \<le> Grp (Collect R) id) = (All R)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   138
  unfolding Grp_def id_apply by blast
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   139
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   140
lemma Grp_id_mono_subst: "(\<And>x y. Grp P id x y \<Longrightarrow> Grp Q id (f x) (f y)) \<equiv>
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   141
   (\<And>x. x \<in> P \<Longrightarrow> f x \<in> Q)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   142
  unfolding Grp_def by rule auto
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   143
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   144
lemma vimage2p_mono: "vimage2p f g R x y \<Longrightarrow> R \<le> S \<Longrightarrow> vimage2p f g S x y"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   145
  unfolding vimage2p_def by blast
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   146
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   147
lemma vimage2p_refl: "(\<And>x. R x x) \<Longrightarrow> vimage2p f f R x x"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   148
  unfolding vimage2p_def by auto
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   149
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   150
lemma
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   151
  assumes "type_definition Rep Abs UNIV"
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   152
  shows type_copy_Rep_o_Abs: "Rep \<circ> Abs = id" and type_copy_Abs_o_Rep: "Abs \<circ> Rep = id"
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   153
  unfolding fun_eq_iff comp_apply id_apply
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   154
    type_definition.Abs_inverse[OF assms UNIV_I] type_definition.Rep_inverse[OF assms] by simp_all
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   155
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   156
lemma type_copy_map_comp0_undo:
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   157
  assumes "type_definition Rep Abs UNIV"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   158
          "type_definition Rep' Abs' UNIV"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   159
          "type_definition Rep'' Abs'' UNIV"
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   160
  shows "Abs' \<circ> M \<circ> Rep'' = (Abs' \<circ> M1 \<circ> Rep) \<circ> (Abs \<circ> M2 \<circ> Rep'') \<Longrightarrow> M1 \<circ> M2 = M"
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   161
  by (rule sym) (auto simp: fun_eq_iff type_definition.Abs_inject[OF assms(2) UNIV_I UNIV_I]
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   162
    type_definition.Abs_inverse[OF assms(1) UNIV_I]
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   163
    type_definition.Abs_inverse[OF assms(3) UNIV_I] dest: spec[of _ "Abs'' x" for x])
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   164
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   165
lemma vimage2p_id: "vimage2p id id R = R"
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   166
  unfolding vimage2p_def by auto
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   167
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   168
lemma vimage2p_comp: "vimage2p (f1 \<circ> f2) (g1 \<circ> g2) = vimage2p f2 g2 \<circ> vimage2p f1 g1"
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   169
  unfolding fun_eq_iff vimage2p_def o_apply by simp
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   170
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   171
lemma vimage2p_rel_fun: "rel_fun (vimage2p f g R) R f g"
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   172
  unfolding rel_fun_def vimage2p_def by auto
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   173
56650
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   174
lemma fun_cong_unused_0: "f = (\<lambda>x. g) \<Longrightarrow> f (\<lambda>x. 0) = g"
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   175
  by (erule arg_cong)
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   176
56684
d8f32f55e463 more unfolding and more folding in size equations, to look more natural in the nested case
blanchet
parents: 56651
diff changeset
   177
lemma inj_on_convol_ident: "inj_on (\<lambda>x. (x, f x)) X"
56650
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   178
  unfolding inj_on_def by simp
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   179
58734
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   180
lemma map_sum_if_distrib_then:
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   181
  "\<And>f g e x y. map_sum f g (if e then Inl x else y) = (if e then Inl (f x) else map_sum f g y)"
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   182
  "\<And>f g e x y. map_sum f g (if e then Inr x else y) = (if e then Inr (g x) else map_sum f g y)"
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   183
  by simp_all
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   184
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   185
lemma map_sum_if_distrib_else:
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   186
  "\<And>f g e x y. map_sum f g (if e then x else Inl y) = (if e then map_sum f g x else Inl (f y))"
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   187
  "\<And>f g e x y. map_sum f g (if e then x else Inr y) = (if e then map_sum f g x else Inr (g y))"
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   188
  by simp_all
20235f0512e2 generate 'map_o_corec' for (co)datatypes
desharna
parents: 58732
diff changeset
   189
56650
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   190
lemma case_prod_app: "case_prod f x y = case_prod (\<lambda>l r. f l r y) x"
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   191
  by (case_tac x) simp
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   192
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   193
lemma case_sum_map_sum: "case_sum l r (map_sum f g x) = case_sum (l \<circ> f) (r \<circ> g) x"
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   194
  by (case_tac x) simp+
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   195
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   196
lemma case_sum_transfer:
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   197
  "rel_fun (rel_fun R T) (rel_fun (rel_fun S T) (rel_fun (rel_sum R S) T)) case_sum case_sum"
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58889
diff changeset
   198
  unfolding rel_fun_def by (auto split: sum.splits)
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   199
56650
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   200
lemma case_prod_map_prod: "case_prod h (map_prod f g x) = case_prod (\<lambda>l r. h (f l) (g r)) x"
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   201
  by (case_tac x) simp+
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   202
58732
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58448
diff changeset
   203
lemma case_prod_o_map_prod: "case_prod f \<circ> map_prod g1 g2 = case_prod (\<lambda>l r. f (g1 l) (g2 r))"
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58448
diff changeset
   204
  unfolding comp_def by auto
854eed6e5aed move theorem 'rec_o_map'
desharna
parents: 58448
diff changeset
   205
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   206
lemma case_prod_transfer:
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   207
  "(rel_fun (rel_fun A (rel_fun B C)) (rel_fun (rel_prod A B) C)) case_prod case_prod"
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58889
diff changeset
   208
  unfolding rel_fun_def by simp
56650
1f9ab71d43a5 no need to make 'size' generation an interpretation -- overkill
blanchet
parents: 56640
diff changeset
   209
57489
8f0ba9f2d10f generate 'corec_code' theorem for codatatypes
desharna
parents: 57471
diff changeset
   210
lemma eq_ifI: "(P \<longrightarrow> t = u1) \<Longrightarrow> (\<not> P \<longrightarrow> t = u2) \<Longrightarrow> t = (if P then u1 else u2)"
8f0ba9f2d10f generate 'corec_code' theorem for codatatypes
desharna
parents: 57471
diff changeset
   211
  by simp
8f0ba9f2d10f generate 'corec_code' theorem for codatatypes
desharna
parents: 57471
diff changeset
   212
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   213
lemma comp_transfer:
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   214
  "rel_fun (rel_fun B C) (rel_fun (rel_fun A B) (rel_fun A C)) (op \<circ>) (op \<circ>)"
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   215
  unfolding rel_fun_def by simp
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58352
diff changeset
   216
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   217
lemma If_transfer: "rel_fun (op =) (rel_fun A (rel_fun A A)) If If"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   218
  unfolding rel_fun_def by simp
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   219
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   220
lemma Abs_transfer:
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   221
  assumes type_copy1: "type_definition Rep1 Abs1 UNIV"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   222
  assumes type_copy2: "type_definition Rep2 Abs2 UNIV"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   223
  shows "rel_fun R (vimage2p Rep1 Rep2 R) Abs1 Abs2"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   224
  unfolding vimage2p_def rel_fun_def
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   225
    type_definition.Abs_inverse[OF type_copy1 UNIV_I]
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   226
    type_definition.Abs_inverse[OF type_copy2 UNIV_I] by simp
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   227
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   228
lemma Inl_transfer:
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   229
  "rel_fun S (rel_sum S T) Inl Inl"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   230
  by auto
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   231
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   232
lemma Inr_transfer:
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   233
  "rel_fun T (rel_sum S T) Inr Inr"
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   234
  by auto
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   235
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   236
lemma Pair_transfer: "rel_fun A (rel_fun B (rel_prod A B)) Pair Pair"
58916
229765cc3414 more complete fp_sugars for sum and prod;
traytel
parents: 58889
diff changeset
   237
  unfolding rel_fun_def by simp
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
   238
62335
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   239
lemma eq_onp_live_step: "x = y \<Longrightarrow> eq_onp P a a \<and> x \<longleftrightarrow> P a \<and> y"
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   240
  by (simp only: eq_onp_same_args)
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   241
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   242
lemma top_conj: "top x \<and> P \<longleftrightarrow> P" "P \<and> top x \<longleftrightarrow> P"
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   243
  by blast+
e85c42f4f30a making 'pred_inject' a first-class BNF citizen
blanchet
parents: 62325
diff changeset
   244
62905
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   245
lemma fst_convol': "fst (\<langle>f, g\<rangle> x) = f x"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   246
  using fst_convol unfolding convol_def by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   247
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   248
lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   249
  using snd_convol unfolding convol_def by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   250
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   251
lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   252
  unfolding convol_def by auto
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   253
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   254
lemma convol_expand_snd':
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   255
  assumes "(fst o f = g)"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   256
  shows "h = snd o f \<longleftrightarrow> \<langle>g, h\<rangle> = f"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   257
proof -
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   258
  from assms have *: "\<langle>g, snd o f\<rangle> = f" by (rule convol_expand_snd)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   259
  then have "h = snd o f \<longleftrightarrow> h = snd o \<langle>g, snd o f\<rangle>" by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   260
  moreover have "\<dots> \<longleftrightarrow> h = snd o f" by (simp add: snd_convol)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   261
  moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   262
  ultimately show ?thesis by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   263
qed
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   264
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   265
lemma case_sum_expand_Inr_pointfree: "f o Inl = g \<Longrightarrow> case_sum g (f o Inr) = f"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   266
  by (auto split: sum.splits)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   267
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   268
lemma case_sum_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> case_sum g h = f"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   269
  by (rule iffI) (auto simp add: fun_eq_iff split: sum.splits)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   270
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   271
lemma case_sum_expand_Inr: "f o Inl = g \<Longrightarrow> f x = case_sum g (f o Inr) x"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   272
  by (auto split: sum.splits)
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   273
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   274
lemma id_transfer: "rel_fun A A id id"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   275
  unfolding rel_fun_def by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   276
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   277
lemma fst_transfer: "rel_fun (rel_prod A B) A fst fst"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   278
  unfolding rel_fun_def by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   279
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   280
lemma snd_transfer: "rel_fun (rel_prod A B) B snd snd"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   281
  unfolding rel_fun_def by simp
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   282
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   283
lemma convol_transfer:
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   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"
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   285
  unfolding rel_fun_def convol_def by auto
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   286
63046
8053ef5a0174 intermediate definitions and caching in n2m to keep terms small
traytel
parents: 62905
diff changeset
   287
lemma Let_const: "Let x (\<lambda>_. c) = c"
8053ef5a0174 intermediate definitions and caching in n2m to keep terms small
traytel
parents: 62905
diff changeset
   288
  unfolding Let_def ..
8053ef5a0174 intermediate definitions and caching in n2m to keep terms small
traytel
parents: 62905
diff changeset
   289
62905
52c5a25e0c96 derive (co)rec uniformly from (un)fold
traytel
parents: 62335
diff changeset
   290
ML_file "Tools/BNF/bnf_fp_util_tactics.ML"
55062
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   291
ML_file "Tools/BNF/bnf_fp_util.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   292
ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   293
ML_file "Tools/BNF/bnf_fp_def_sugar.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   294
ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   295
ML_file "Tools/BNF/bnf_fp_n2m.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   296
ML_file "Tools/BNF/bnf_fp_n2m_sugar.ML"
55702
63c80031d8dd improved accounting for dead variables when naming set functions (refines d71c2737ee21)
blanchet
parents: 55700
diff changeset
   297
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
   298
end