src/HOL/BNF_FP_Base.thy
author blanchet
Wed, 23 Apr 2014 10:23:26 +0200
changeset 56640 0a35354137a5
parent 55966 972f0aa7091b
child 56650 1f9ab71d43a5
permissions -rw-r--r--
generate 'rec_o_map' and 'size_o_map' in size extension
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55059
ef2e0fb783c6 tuned comments
blanchet
parents: 55058
diff changeset
     1
(*  Title:      HOL/BNF_FP_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
53311
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
     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
ef2e0fb783c6 tuned comments
blanchet
parents: 55058
diff changeset
     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
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
    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
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
    12
theory BNF_FP_Base
55936
f6591f8c629d rationalized imports
traytel
parents: 55932
diff changeset
    13
imports BNF_Comp 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
    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
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    16
lemma mp_conj: "(P \<longrightarrow> Q) \<and> R \<Longrightarrow> P \<Longrightarrow> R \<and> Q"
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    17
by auto
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    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
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    22
lemma case_unit_Unity: "(case u of () \<Rightarrow> f) = f"
55642
63beb38e9258 adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
blanchet
parents: 55575
diff changeset
    23
by (cases u) (hypsubst, rule unit.case)
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    24
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    25
lemma case_prod_Pair_iden: "(case p of (x, y) \<Rightarrow> (x, y)) = p"
49539
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
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    28
lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    29
by simp
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    30
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    31
lemma pointfree_idE: "f \<circ> g = id \<Longrightarrow> f (g x) = x"
55066
blanchet
parents: 55062
diff changeset
    32
unfolding comp_def fun_eq_iff by simp
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    33
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    34
lemma o_bij:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    35
  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
    36
  shows "bij f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    37
unfolding bij_def inj_on_def surj_def proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    38
  fix a1 a2 assume "f a1 = f a2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    39
  hence "g ( f a1) = g (f a2)" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    40
  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
    41
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    42
  fix b
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
  have "b = f (g b)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    44
  using fg unfolding fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    45
  thus "EX a. b = f a" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    49
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    50
lemma case_sum_step:
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    51
"case_sum (case_sum f' g') g (Inl p) = case_sum f' g' p"
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    52
"case_sum f (case_sum f' g') (Inr p) = case_sum f' g' p"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    53
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    54
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    55
lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    56
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    57
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
    58
lemma type_copy_obj_one_point_absE:
55811
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    59
  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
    60
  using type_definition.Rep_inverse[OF assms(1)]
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    61
  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
    62
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    63
lemma obj_sumE_f:
55811
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    64
  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
    65
  shows "\<forall>x. s = f x \<longrightarrow> P"
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    66
proof
aa1acc25126b load Metis a little later
traytel
parents: 55803
diff changeset
    67
  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
    68
qed
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    69
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    70
lemma case_sum_if:
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55083
diff changeset
    71
"case_sum f g (if p then Inl x else Inr y) = (if p then f x else g y)"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    72
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    73
51745
a06a3c777add simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents: 51740
diff changeset
    74
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
    75
by simp
a06a3c777add simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents: 51740
diff changeset
    76
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    77
lemma prod_set_simps:
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    78
"fsts (x, y) = {x}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    79
"snds (x, y) = {y}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    80
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
    81
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    82
lemma sum_set_simps:
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    83
"setl (Inl x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    84
"setl (Inr x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    85
"setr (Inl x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    86
"setr (Inr x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    87
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
    88
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
    89
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
    90
by blast
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
    91
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
    92
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
    93
  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
    94
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
    95
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
    96
  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
    97
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
    98
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
    99
  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
   100
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   101
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
   102
  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
   103
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   104
lemma convol_o: "<f, g> \<circ> h = <f \<circ> h, g \<circ> h>"
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   105
  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
   106
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   107
lemma map_prod_o_convol: "map_prod h1 h2 \<circ> <f, g> = <h1 \<circ> f, h2 \<circ> g>"
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   108
  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
   109
55932
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55931
diff changeset
   110
lemma map_prod_o_convol_id: "(map_prod f id \<circ> <id , g>) x = <id \<circ> f , g> x"
68c5104d2204 renamed 'map_pair' to 'map_prod'
blanchet
parents: 55931
diff changeset
   111
  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
   112
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   113
lemma o_case_sum: "h \<circ> case_sum f g = case_sum (h \<circ> f) (h \<circ> g)"
55066
blanchet
parents: 55062
diff changeset
   114
  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
   115
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   116
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
   117
  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
   118
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   119
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
   120
  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
   121
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55936
diff changeset
   122
lemma rel_fun_def_butlast:
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55936
diff changeset
   123
  "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
   124
  unfolding rel_fun_def ..
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   125
53105
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   126
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
   127
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   128
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   129
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
   130
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   131
53308
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   132
lemma eq_le_Grp_id_iff: "(op = \<le> Grp (Collect R) id) = (All R)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   133
  unfolding Grp_def id_apply by blast
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   134
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   135
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
   136
   (\<And>x. x \<in> P \<Longrightarrow> f x \<in> Q)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   137
  unfolding Grp_def by rule auto
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   138
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   139
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
   140
  unfolding vimage2p_def by blast
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   141
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   142
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
   143
  unfolding vimage2p_def by auto
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   144
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   145
lemma
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   146
  assumes "type_definition Rep Abs UNIV"
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   147
  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
   148
  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
   149
    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
   150
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   151
lemma type_copy_map_comp0_undo:
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   152
  assumes "type_definition Rep Abs UNIV"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   153
          "type_definition Rep' Abs' UNIV"
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55702
diff changeset
   154
          "type_definition Rep'' Abs'' UNIV"
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   155
  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
   156
  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
   157
    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
   158
    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
   159
55854
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   160
lemma vimage2p_id: "vimage2p id id R = R"
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   161
  unfolding vimage2p_def by auto
ee270328a781 make 'typedef' optional, depending on size of original type
blanchet
parents: 55811
diff changeset
   162
56640
0a35354137a5 generate 'rec_o_map' and 'size_o_map' in size extension
blanchet
parents: 55966
diff changeset
   163
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
   164
  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
   165
55062
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   166
ML_file "Tools/BNF/bnf_fp_util.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   167
ML_file "Tools/BNF/bnf_fp_def_sugar_tactics.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   168
ML_file "Tools/BNF/bnf_fp_def_sugar.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   169
ML_file "Tools/BNF/bnf_fp_n2m_tactics.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   170
ML_file "Tools/BNF/bnf_fp_n2m.ML"
6d3fad6f01c9 made BNF compile after move to HOL
blanchet
parents: 55059
diff changeset
   171
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
   172
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
   173
end