src/HOL/BNF/BNF_FP_Base.thy
author blanchet
Thu, 26 Sep 2013 01:05:06 +0200
changeset 53900 527ece7edc51
parent 53868 c25acff63bfe
child 53903 27fd72593624
permissions -rw-r--r--
made tactic more flexible w.r.t. case expressions and such
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53311
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
     1
(*  Title:      HOL/BNF/BNF_FP_Base.thy
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
53311
802ae7dae691 tuned theory name
blanchet
parents: 53310
diff changeset
     7
Shared fixed point operations on bounded natural functors, including
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
51797
182454c06a80 tuned ML and thy file names
blanchet
parents: 51745
diff changeset
    13
imports BNF_Comp BNF_Ctr_Sugar
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
53900
527ece7edc51 made tactic more flexible w.r.t. case expressions and such
blanchet
parents: 53868
diff changeset
    16
lemma not_TrueE: "\<not> True \<Longrightarrow> P"
527ece7edc51 made tactic more flexible w.r.t. case expressions and such
blanchet
parents: 53868
diff changeset
    17
by (erule notE, rule TrueI)
527ece7edc51 made tactic more flexible w.r.t. case expressions and such
blanchet
parents: 53868
diff changeset
    18
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    19
lemma mp_conj: "(P \<longrightarrow> Q) \<and> R \<Longrightarrow> P \<Longrightarrow> R \<and> Q"
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    20
by auto
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    21
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    22
lemma eq_sym_Unity_conv: "(x = (() = ())) = x"
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    23
by blast
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    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 unit_case_Unity: "(case u of () => f) = f"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    26
by (cases u) (hypsubst, rule unit.cases)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    27
49539
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    28
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
    29
by simp
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    30
49335
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    31
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
    32
by simp
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    33
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    34
lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    35
by clarify
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    36
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    37
lemma prod_all_impI_step: "(\<And>x. \<forall>y. P (x, y) \<longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    38
by auto
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    39
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    40
lemma pointfree_idE: "f \<circ> g = id \<Longrightarrow> f (g x) = x"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    41
unfolding o_def fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    42
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
lemma o_bij:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    44
  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
    45
  shows "bij f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
unfolding bij_def inj_on_def surj_def proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
  fix a1 a2 assume "f a1 = f a2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
  hence "g ( f a1) = g (f a2)" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    49
  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
    50
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    51
  fix b
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    52
  have "b = f (g b)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    53
  using fg unfolding fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    54
  thus "EX a. b = f a" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    55
qed
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 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
    58
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    59
lemma sum_case_step:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    60
"sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    61
"sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    62
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    63
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    64
lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    65
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    66
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    67
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
    68
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    69
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    70
lemma obj_sumE_f:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    71
"\<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
7f7311d04727 killed unused theorems
traytel
parents: 52659
diff changeset
    72
by (rule allI) (metis sumE)
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    73
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    74
lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    75
by (cases s) auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    76
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    77
lemma sum_case_if:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    78
"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    79
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    80
49428
93f158d59ead handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents: 49427
diff changeset
    81
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
    82
by blast
93f158d59ead handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents: 49427
diff changeset
    83
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    84
lemma UN_compreh_eq_eq:
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    85
"\<Union>{y. \<exists>x\<in>A. y = {}} = {}"
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    86
"\<Union>{y. \<exists>x\<in>A. y = {x}} = A"
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    87
by blast+
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    88
51745
a06a3c777add simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents: 51740
diff changeset
    89
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
    90
by simp
a06a3c777add simplify "Inl () = Inr ()" as well (not entirely clear why this is necessary)
blanchet
parents: 51740
diff changeset
    91
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    92
lemma prod_set_simps:
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    93
"fsts (x, y) = {x}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    94
"snds (x, y) = {y}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    95
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
    96
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
    97
lemma sum_set_simps:
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    98
"setl (Inl x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
    99
"setl (Inr x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   100
"setr (Inl x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   101
"setr (Inr x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   102
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
   103
49642
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   104
lemma prod_rel_simp:
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   105
"prod_rel P Q (x, y) (x', y') \<longleftrightarrow> P x x' \<and> Q y y'"
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   106
unfolding prod_rel_def by simp
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   107
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   108
lemma sum_rel_simps:
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   109
"sum_rel P Q (Inl x) (Inl x') \<longleftrightarrow> P x x'"
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   110
"sum_rel P Q (Inr y) (Inr y') \<longleftrightarrow> Q y y'"
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   111
"sum_rel P Q (Inl x) (Inr y') \<longleftrightarrow> False"
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   112
"sum_rel P Q (Inr y) (Inl x') \<longleftrightarrow> False"
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   113
unfolding sum_rel_def by simp+
9f884142334c fixed simplification of prod and sum relators to avoid issues with e.g. codata ('a, 'b) k = K "'a + 'b"
blanchet
parents: 49636
diff changeset
   114
52505
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
   115
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
   116
by blast
e62f3fd2035e share some code between codatatypes, datatypes and eventually prim(co)rec
traytel
parents: 52349
diff changeset
   117
52913
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   118
lemma rewriteR_comp_comp: "\<lbrakk>g o h = r\<rbrakk> \<Longrightarrow> f o g o h = f o r"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   119
  unfolding o_def fun_eq_iff by auto
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   120
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   121
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"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   122
  unfolding o_def fun_eq_iff by auto
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   123
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   124
lemma rewriteL_comp_comp: "\<lbrakk>f o g = l\<rbrakk> \<Longrightarrow> f o (g o h) = l o h"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   125
  unfolding o_def fun_eq_iff by auto
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   126
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   127
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"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   128
  unfolding o_def fun_eq_iff by auto
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   129
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   130
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
   131
  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
   132
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   133
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
   134
  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
   135
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   136
lemma map_pair_o_convol_id: "(map_pair f id \<circ> <id , g>) x = <id \<circ> f , g> x"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   137
  unfolding map_pair_o_convol id_o o_id ..
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   138
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   139
lemma o_sum_case: "h o sum_case f g = sum_case (h o f) (h o g)"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   140
  unfolding o_def by (auto split: sum.splits)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   141
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   142
lemma sum_case_o_sum_map: "sum_case f g o sum_map h1 h2 = sum_case (f o h1) (g o h2)"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   143
  unfolding o_def by (auto split: sum.splits)
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   144
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   145
lemma sum_case_o_sum_map_id: "(sum_case id g o sum_map f id) x = sum_case (f o id) g x"
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   146
  unfolding sum_case_o_sum_map id_o o_id ..
2d2d9d1de1a9 theorems relating {c,d}tor_(un)fold/(co)rec and {c,d}tor_map
traytel
parents: 52731
diff changeset
   147
52731
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   148
lemma fun_rel_def_butlast:
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   149
  "(fun_rel R (fun_rel S T)) f g = (\<forall>x y. R x y \<longrightarrow> (fun_rel S T) (f x) (g y))"
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   150
  unfolding fun_rel_def ..
dacd47a0633f transfer rule for {c,d}tor_{,un}fold
traytel
parents: 52660
diff changeset
   151
53105
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   152
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
   153
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   154
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   155
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
   156
  by auto
ec38e9f4352f simpler (forward) derivation of strong (up-to equality) coinduction properties
traytel
parents: 52913
diff changeset
   157
53308
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   158
lemma eq_le_Grp_id_iff: "(op = \<le> Grp (Collect R) id) = (All R)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   159
  unfolding Grp_def id_apply by blast
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   160
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   161
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
   162
   (\<And>x. x \<in> P \<Longrightarrow> f x \<in> Q)"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   163
  unfolding Grp_def by rule auto
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   164
53693
71b020d161c5 new tactics for constructor view
blanchet
parents: 53553
diff changeset
   165
lemma eq_ifI: "\<lbrakk>b \<Longrightarrow> t = x; \<not> b \<Longrightarrow> t = y\<rbrakk> \<Longrightarrow> t = (if b then x else y)"
71b020d161c5 new tactics for constructor view
blanchet
parents: 53553
diff changeset
   166
  by fastforce
71b020d161c5 new tactics for constructor view
blanchet
parents: 53553
diff changeset
   167
51850
106afdf5806c renamed a few FP-related files, to make it clear that these are not the sum of LFP + GFP but rather shared basic libraries
blanchet
parents: 51797
diff changeset
   168
ML_file "Tools/bnf_fp_util.ML"
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
   169
ML_file "Tools/bnf_fp_def_sugar_tactics.ML"
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
   170
ML_file "Tools/bnf_fp_def_sugar.ML"
53308
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   171
ML_file "Tools/bnf_fp_n2m_tactics.ML"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   172
ML_file "Tools/bnf_fp_n2m.ML"
d066e4923a31 merged two theory files
blanchet
parents: 53305
diff changeset
   173
ML_file "Tools/bnf_fp_n2m_sugar.ML"
53305
29c267cb9314 rationalized files
blanchet
parents: 53105
diff changeset
   174
ML_file "Tools/bnf_fp_rec_sugar_util.ML"
29c267cb9314 rationalized files
blanchet
parents: 53105
diff changeset
   175
ML_file "Tools/bnf_fp_rec_sugar_tactics.ML"
29c267cb9314 rationalized files
blanchet
parents: 53105
diff changeset
   176
ML_file "Tools/bnf_fp_rec_sugar.ML"
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   177
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
   178
end