src/HOL/BNF/BNF_FP.thy
author blanchet
Tue, 02 Oct 2012 01:00:18 +0200
changeset 49683 78a3d5006cf1
parent 49642 9f884142334c
child 51740 97c116445b65
permissions -rw-r--r--
continued changing type of corec type
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49495
diff changeset
     1
(*  Title:      HOL/BNF/BNF_FP.thy
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
     2
    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
     3
    Author:     Jasmin Blanchette, 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
    Copyright   2012
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
     5
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
     6
Composition of bounded natural functors.
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
     7
*)
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
header {* Composition of Bounded Natural Functors *}
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    10
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    11
theory BNF_FP
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    12
imports BNF_Comp BNF_Wrap
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    13
keywords
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
    14
  "defaults"
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
49590
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    17
lemma mp_conj: "(P \<longrightarrow> Q) \<and> R \<Longrightarrow> P \<Longrightarrow> R \<and> Q"
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    18
by auto
43e2d0e10876 added coinduction tactic
blanchet
parents: 49585
diff changeset
    19
49591
91b228e26348 generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents: 49590
diff changeset
    20
lemma eq_sym_Unity_conv: "(x = (() = ())) = x"
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    21
by blast
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
    22
49539
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    23
lemma unit_case_Unity: "(case u of () => f) = f"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    24
by (cases u) (hypsubst, rule unit.cases)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    25
49539
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    26
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
    27
by simp
be6cbf960aa7 fixed bug in "fold" tactic with nested products (beyond the sum of product corresponding to constructors)
blanchet
parents: 49510
diff changeset
    28
49335
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    29
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
    30
by simp
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    31
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    32
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
    33
by clarify
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    34
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    35
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
    36
by auto
096967bf3940 added sumEN_tupled_balanced
blanchet
parents: 49325
diff changeset
    37
49368
df359ce891ac added induct tactic
blanchet
parents: 49335
diff changeset
    38
lemma all_unit_eq: "(\<And>x. PROP P x) \<equiv> PROP P ()"
df359ce891ac added induct tactic
blanchet
parents: 49335
diff changeset
    39
by simp
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    40
49368
df359ce891ac added induct tactic
blanchet
parents: 49335
diff changeset
    41
lemma all_prod_eq: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))"
df359ce891ac added induct tactic
blanchet
parents: 49335
diff changeset
    42
by clarsimp
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    44
lemma rev_bspec: "a \<in> A \<Longrightarrow> \<forall>z \<in> A. P z \<Longrightarrow> P a"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    45
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
lemma Un_cong: "\<lbrakk>A = B; C = D\<rbrakk> \<Longrightarrow> A \<union> C = B \<union> D"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    49
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    50
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
    51
unfolding o_def fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    52
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    53
lemma o_bij:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    54
  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
    55
  shows "bij f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    56
unfolding bij_def inj_on_def surj_def proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    57
  fix a1 a2 assume "f a1 = f a2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    58
  hence "g ( f a1) = g (f a2)" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    59
  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
    60
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    61
  fix b
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    62
  have "b = f (g b)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    63
  using fg unfolding fun_eq_iff by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    64
  thus "EX a. b = f a" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    65
qed
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 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
    68
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    69
lemma sum_case_step:
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
    70
"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
    71
"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
    72
by auto
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 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
    75
by simp
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 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
    78
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    79
49325
blanchet
parents: 49312
diff changeset
    80
lemma obj_sumE_f':
blanchet
parents: 49312
diff changeset
    81
"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f x \<longrightarrow> P"
blanchet
parents: 49312
diff changeset
    82
by (cases x) blast+
blanchet
parents: 49312
diff changeset
    83
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    84
lemma obj_sumE_f:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    85
"\<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"
49325
blanchet
parents: 49312
diff changeset
    86
by (rule allI) (rule obj_sumE_f')
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    87
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    88
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
    89
by (cases s) auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    90
49325
blanchet
parents: 49312
diff changeset
    91
lemma obj_sum_step':
blanchet
parents: 49312
diff changeset
    92
"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f (Inr x) \<longrightarrow> P"
blanchet
parents: 49312
diff changeset
    93
by (cases x) blast+
blanchet
parents: 49312
diff changeset
    94
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    95
lemma obj_sum_step:
49325
blanchet
parents: 49312
diff changeset
    96
"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f (Inr x) \<longrightarrow> P"
blanchet
parents: 49312
diff changeset
    97
by (rule allI) (rule obj_sum_step')
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    98
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    99
lemma sum_case_if:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   100
"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
   101
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   102
49683
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   103
lemma sum_case_o_inj:
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   104
"sum_case f g \<circ> Inl = f"
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   105
"sum_case f g \<circ> Inr = g"
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   106
by auto
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   107
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   108
lemma ident_o_ident: "(\<lambda>x. x) \<circ> (\<lambda>x. x) = (\<lambda>x. x)"
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   109
by (rule o_def)
78a3d5006cf1 continued changing type of corec type
blanchet
parents: 49642
diff changeset
   110
49428
93f158d59ead handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents: 49427
diff changeset
   111
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
   112
by blast
93f158d59ead handle the general case with more than two levels of nesting when discharging induction prem prems
blanchet
parents: 49427
diff changeset
   113
49585
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
   114
lemma UN_compreh_eq_eq:
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
   115
"\<Union>{y. \<exists>x\<in>A. y = {}} = {}"
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
   116
"\<Union>{y. \<exists>x\<in>A. y = {x}} = A"
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
   117
by blast+
5c4a12550491 generate high-level "maps", "sets", and "rels" properties
blanchet
parents: 49539
diff changeset
   118
49429
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   119
lemma prod_set_simps:
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   120
"fsts (x, y) = {x}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   121
"snds (x, y) = {y}"
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   122
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
   123
64ac3471005a cleaner way of dealing with the set functions of sums and products
blanchet
parents: 49428
diff changeset
   124
lemma sum_set_simps:
49451
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   125
"setl (Inl x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   126
"setl (Inr x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   127
"setr (Inl x) = {}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   128
"setr (Inr x) = {x}"
7a28d22c33c6 renamed "sum_setl" to "setl" and similarly for r
blanchet
parents: 49430
diff changeset
   129
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
   130
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
   131
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
   132
"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
   133
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
   134
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
   135
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
   136
"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
   137
"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
   138
"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
   139
"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
   140
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
   141
49457
1d2825673cec renamed "bnf_fp_util.ML" to "bnf_fp.ML"
blanchet
parents: 49451
diff changeset
   142
ML_file "Tools/bnf_fp.ML"
49636
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
   143
ML_file "Tools/bnf_fp_def_sugar_tactics.ML"
b7256a88a84b renamed ML file in preparation for next step
blanchet
parents: 49591
diff changeset
   144
ML_file "Tools/bnf_fp_def_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
   145
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents:
diff changeset
   146
end