src/HOL/Codatatype/BNF_FP.thy
author blanchet
Mon, 17 Sep 2012 21:13:30 +0200
changeset 50444 64ac3471005a
parent 50443 93f158d59ead
child 50445 6df729c6a1a6
permissions -rw-r--r--
cleaner way of dealing with the set functions of sums and products
blanchet@50323
     1
(*  Title:      HOL/Codatatype/BNF_FP.thy
blanchet@50323
     2
    Author:     Dmitriy Traytel, TU Muenchen
blanchet@50323
     3
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@50323
     4
    Copyright   2012
blanchet@50323
     5
blanchet@50323
     6
Composition of bounded natural functors.
blanchet@50323
     7
*)
blanchet@50323
     8
blanchet@50323
     9
header {* Composition of Bounded Natural Functors *}
blanchet@50323
    10
blanchet@50323
    11
theory BNF_FP
blanchet@50323
    12
imports BNF_Comp BNF_Wrap
blanchet@50323
    13
keywords
blanchet@50323
    14
  "defaults"
blanchet@50323
    15
begin
blanchet@50323
    16
blanchet@50327
    17
lemma case_unit: "(case u of () => f) = f"
blanchet@50327
    18
by (cases u) (hypsubst, rule unit.cases)
blanchet@50327
    19
blanchet@50350
    20
lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
blanchet@50350
    21
by simp
blanchet@50350
    22
blanchet@50350
    23
lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
blanchet@50350
    24
by clarify
blanchet@50350
    25
blanchet@50350
    26
lemma prod_all_impI_step: "(\<And>x. \<forall>y. P (x, y) \<longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
blanchet@50350
    27
by auto
blanchet@50350
    28
blanchet@50383
    29
lemma all_unit_eq: "(\<And>x. PROP P x) \<equiv> PROP P ()"
blanchet@50383
    30
by simp
blanchet@50327
    31
blanchet@50383
    32
lemma all_prod_eq: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))"
blanchet@50383
    33
by clarsimp
blanchet@50327
    34
blanchet@50327
    35
lemma rev_bspec: "a \<in> A \<Longrightarrow> \<forall>z \<in> A. P z \<Longrightarrow> P a"
blanchet@50327
    36
by simp
blanchet@50327
    37
blanchet@50327
    38
lemma Un_cong: "\<lbrakk>A = B; C = D\<rbrakk> \<Longrightarrow> A \<union> C = B \<union> D"
blanchet@50327
    39
by simp
blanchet@50327
    40
blanchet@50327
    41
definition convol ("<_ , _>") where
blanchet@50327
    42
"<f , g> \<equiv> %a. (f a, g a)"
blanchet@50327
    43
blanchet@50327
    44
lemma fst_convol:
blanchet@50327
    45
"fst o <f , g> = f"
blanchet@50327
    46
apply(rule ext)
blanchet@50327
    47
unfolding convol_def by simp
blanchet@50327
    48
blanchet@50327
    49
lemma snd_convol:
blanchet@50327
    50
"snd o <f , g> = g"
blanchet@50327
    51
apply(rule ext)
blanchet@50327
    52
unfolding convol_def by simp
blanchet@50327
    53
blanchet@50327
    54
lemma pointfree_idE: "f o g = id \<Longrightarrow> f (g x) = x"
blanchet@50327
    55
unfolding o_def fun_eq_iff by simp
blanchet@50327
    56
blanchet@50327
    57
lemma o_bij:
blanchet@50327
    58
  assumes gf: "g o f = id" and fg: "f o g = id"
blanchet@50327
    59
  shows "bij f"
blanchet@50327
    60
unfolding bij_def inj_on_def surj_def proof safe
blanchet@50327
    61
  fix a1 a2 assume "f a1 = f a2"
blanchet@50327
    62
  hence "g ( f a1) = g (f a2)" by simp
blanchet@50327
    63
  thus "a1 = a2" using gf unfolding fun_eq_iff by simp
blanchet@50327
    64
next
blanchet@50327
    65
  fix b
blanchet@50327
    66
  have "b = f (g b)"
blanchet@50327
    67
  using fg unfolding fun_eq_iff by simp
blanchet@50327
    68
  thus "EX a. b = f a" by blast
blanchet@50327
    69
qed
blanchet@50327
    70
blanchet@50327
    71
lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
blanchet@50327
    72
blanchet@50327
    73
lemma sum_case_step:
blanchet@50327
    74
  "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
blanchet@50327
    75
  "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
blanchet@50327
    76
by auto
blanchet@50327
    77
blanchet@50327
    78
lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
blanchet@50327
    79
by simp
blanchet@50327
    80
blanchet@50327
    81
lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
blanchet@50327
    82
by blast
blanchet@50327
    83
blanchet@50340
    84
lemma obj_sumE_f':
blanchet@50340
    85
"\<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@50340
    86
by (cases x) blast+
blanchet@50340
    87
blanchet@50327
    88
lemma obj_sumE_f:
blanchet@50327
    89
"\<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"
blanchet@50340
    90
by (rule allI) (rule obj_sumE_f')
blanchet@50327
    91
blanchet@50327
    92
lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
blanchet@50327
    93
by (cases s) auto
blanchet@50327
    94
blanchet@50340
    95
lemma obj_sum_step':
blanchet@50340
    96
"\<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@50340
    97
by (cases x) blast+
blanchet@50340
    98
blanchet@50327
    99
lemma obj_sum_step:
blanchet@50340
   100
"\<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@50340
   101
by (rule allI) (rule obj_sum_step')
blanchet@50327
   102
blanchet@50327
   103
lemma sum_case_if:
blanchet@50327
   104
"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
blanchet@50327
   105
by simp
blanchet@50327
   106
blanchet@50441
   107
lemma UN_compreh_bex_eq_empty:
blanchet@50441
   108
"\<Union>{y. \<exists>x\<in>A. y = {}} = {}"
blanchet@50441
   109
by blast
blanchet@50441
   110
blanchet@50441
   111
lemma UN_compreh_bex_eq_singleton:
blanchet@50441
   112
"\<Union>{y. \<exists>x\<in>A. y = {f x}} = {y. \<exists>x\<in>A. y = f x}"
blanchet@50441
   113
by blast
blanchet@50441
   114
blanchet@50441
   115
lemma mem_UN_comprehI:
blanchet@50441
   116
"z \<in> {y. \<exists>x\<in>A. y = f x} \<Longrightarrow> z \<in> \<Union>{y. \<exists>x\<in>A. y = {f x}}"
blanchet@50441
   117
"z \<in> {y. \<exists>x\<in>A. y = f x} \<union> B \<Longrightarrow> z \<in> \<Union>{y. \<exists>x\<in>A. y = {f x}} \<union> B"
blanchet@50441
   118
"z \<in> \<Union>{y. \<exists>x\<in>A. y = F x} \<union> \<Union>{y. \<exists>x\<in>A. y = G x} \<Longrightarrow> z \<in> \<Union>{y. \<exists>x\<in>A. y = F x \<union> G x}"
blanchet@50441
   119
"z \<in> \<Union>{y. \<exists>x\<in>A. y = F x} \<union> (\<Union>{y. \<exists>x\<in>A. y = G x} \<union> B) \<Longrightarrow> z \<in> \<Union>{y. \<exists>x\<in>A. y = F x \<union> G x} \<union> B"
blanchet@50383
   120
by blast+
blanchet@50383
   121
blanchet@50441
   122
lemma mem_UN_comprehI':
blanchet@50441
   123
"\<exists>y. (\<exists>x\<in>A. y = F x) \<and> z \<in> y \<Longrightarrow> z \<in> \<Union>{y. \<exists>x\<in>A. y = {y. \<exists>y'\<in>F x. y = y'}}"
blanchet@50441
   124
by blast
blanchet@50441
   125
blanchet@50443
   126
lemma mem_UN_compreh_eq: "(z : \<Union>{y. \<exists>x\<in>A. y = F x}) = (\<exists>x\<in>A. z : F x)"
blanchet@50443
   127
by blast
blanchet@50443
   128
blanchet@50443
   129
lemma eq_UN_compreh_Un: "(xa = \<Union>{y. \<exists>x\<in>A. y = c_set1 x \<union> c_set2 x}) =
blanchet@50443
   130
  (xa = (\<Union>{y. \<exists>x\<in>A. y = c_set1 x} \<union> \<Union>{y. \<exists>x\<in>A. y = c_set2 x}))"
blanchet@50443
   131
by blast
blanchet@50443
   132
blanchet@50442
   133
lemma mem_compreh_eq_iff:
blanchet@50443
   134
"z \<in> {y. \<exists>x\<in>A. y = f x} = (\<exists> x. x \<in> A & z \<in> {f x})"
blanchet@50442
   135
by blast
blanchet@50442
   136
blanchet@50442
   137
lemma ex_mem_singleton: "(\<exists>y. y \<in> A \<and> y \<in> {x}) = (x \<in> A)"
blanchet@50442
   138
by simp
blanchet@50442
   139
blanchet@50444
   140
lemma prod_set_simps:
blanchet@50444
   141
"fsts (x, y) = {x}"
blanchet@50444
   142
"snds (x, y) = {y}"
blanchet@50444
   143
unfolding fsts_def snds_def by simp+
blanchet@50444
   144
blanchet@50444
   145
lemma sum_set_simps:
blanchet@50444
   146
"sum_setl (Inl x) = {x}"
blanchet@50444
   147
"sum_setl (Inr x) = {}"
blanchet@50444
   148
"sum_setr (Inl x) = {}"
blanchet@50444
   149
"sum_setr (Inr x) = {x}"
blanchet@50444
   150
unfolding sum_setl_def sum_setr_def by simp+
blanchet@50444
   151
blanchet@50442
   152
lemma induct_set_step:
blanchet@50442
   153
"\<lbrakk>b \<in> A; c \<in> F b\<rbrakk> \<Longrightarrow> \<exists>x. x \<in> A \<and> c \<in> F x"
blanchet@50442
   154
"\<lbrakk>B \<in> A; c \<in> f B\<rbrakk> \<Longrightarrow> \<exists>C. (\<exists>a \<in> A. C = f a) \<and>c \<in> C"
blanchet@50442
   155
by blast+
blanchet@50383
   156
blanchet@50324
   157
ML_file "Tools/bnf_fp_util.ML"
blanchet@50324
   158
ML_file "Tools/bnf_fp_sugar_tactics.ML"
blanchet@50324
   159
ML_file "Tools/bnf_fp_sugar.ML"
blanchet@50324
   160
blanchet@50323
   161
end