src/HOL/Codatatype/BNF_FP.thy
author blanchet
Fri, 14 Sep 2012 22:23:11 +0200
changeset 50398 0f71da2988d2
parent 50387 c62165188971
child 50404 da621dc65146
permissions -rw-r--r--
killed spurious rotate_tac; use auto instead of blast
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@50398
   107
(* TODO: cleanup *)
blanchet@50383
   108
lemma UN_compreh_bex:
blanchet@50383
   109
"\<Union>{y. \<exists>x \<in> A. y = {}} = {}"
blanchet@50383
   110
"\<Union>{y. \<exists>x \<in> A. y = {x}} = A"
blanchet@50387
   111
"\<Union>{y. \<exists>x \<in> A. y = {f x}} = {y. \<exists>x \<in> A. y = f x}"
blanchet@50383
   112
by blast+
blanchet@50383
   113
blanchet@50383
   114
lemma induct_set_step: "\<lbrakk>B \<in> A; c \<in> f B\<rbrakk> \<Longrightarrow> \<exists>C. (\<exists>a \<in> A. C = f a) \<and> c \<in> C"
blanchet@50383
   115
apply (rule exI)
blanchet@50383
   116
apply (rule conjI)
blanchet@50383
   117
 apply (rule bexI)
blanchet@50383
   118
  apply (rule refl)
blanchet@50383
   119
 apply assumption
blanchet@50383
   120
apply assumption
blanchet@50383
   121
done
blanchet@50383
   122
blanchet@50324
   123
ML_file "Tools/bnf_fp_util.ML"
blanchet@50324
   124
ML_file "Tools/bnf_fp_sugar_tactics.ML"
blanchet@50324
   125
ML_file "Tools/bnf_fp_sugar.ML"
blanchet@50324
   126
blanchet@50323
   127
end