src/HOL/Codatatype/BNF_FP.thy
author blanchet
Wed, 12 Sep 2012 17:26:05 +0200
changeset 50350 096967bf3940
parent 50340 340844cbf7af
child 50383 df359ce891ac
permissions -rw-r--r--
added sumEN_tupled_balanced
     1 (*  Title:      HOL/Codatatype/BNF_FP.thy
     2     Author:     Dmitriy Traytel, TU Muenchen
     3     Author:     Jasmin Blanchette, TU Muenchen
     4     Copyright   2012
     5 
     6 Composition of bounded natural functors.
     7 *)
     8 
     9 header {* Composition of Bounded Natural Functors *}
    10 
    11 theory BNF_FP
    12 imports BNF_Comp BNF_Wrap
    13 keywords
    14   "defaults"
    15 begin
    16 
    17 lemma case_unit: "(case u of () => f) = f"
    18 by (cases u) (hypsubst, rule unit.cases)
    19 
    20 lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
    21 by simp
    22 
    23 lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
    24 by clarify
    25 
    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"
    27 by auto
    28 
    29 lemma all_unit_eq: "(\<And>x. PROP P x) \<equiv> PROP P ()" by simp
    30 
    31 lemma all_prod_eq: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))" by clarsimp
    32 
    33 (* FIXME: needed? *)
    34 lemma False_imp_eq: "(False \<Longrightarrow> P) \<equiv> Trueprop True"
    35 by presburger
    36 
    37 (* FIXME: needed? *)
    38 lemma all_point_1: "(\<And>z. z = b \<Longrightarrow> phi z) \<equiv> Trueprop (phi b)"
    39 by presburger
    40 
    41 lemma rev_bspec: "a \<in> A \<Longrightarrow> \<forall>z \<in> A. P z \<Longrightarrow> P a"
    42 by simp
    43 
    44 lemma Un_cong: "\<lbrakk>A = B; C = D\<rbrakk> \<Longrightarrow> A \<union> C = B \<union> D"
    45 by simp
    46 
    47 definition convol ("<_ , _>") where
    48 "<f , g> \<equiv> %a. (f a, g a)"
    49 
    50 lemma fst_convol:
    51 "fst o <f , g> = f"
    52 apply(rule ext)
    53 unfolding convol_def by simp
    54 
    55 lemma snd_convol:
    56 "snd o <f , g> = g"
    57 apply(rule ext)
    58 unfolding convol_def by simp
    59 
    60 lemma pointfree_idE: "f o g = id \<Longrightarrow> f (g x) = x"
    61 unfolding o_def fun_eq_iff by simp
    62 
    63 lemma o_bij:
    64   assumes gf: "g o f = id" and fg: "f o g = id"
    65   shows "bij f"
    66 unfolding bij_def inj_on_def surj_def proof safe
    67   fix a1 a2 assume "f a1 = f a2"
    68   hence "g ( f a1) = g (f a2)" by simp
    69   thus "a1 = a2" using gf unfolding fun_eq_iff by simp
    70 next
    71   fix b
    72   have "b = f (g b)"
    73   using fg unfolding fun_eq_iff by simp
    74   thus "EX a. b = f a" by blast
    75 qed
    76 
    77 lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
    78 
    79 lemma sum_case_step:
    80   "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
    81   "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
    82 by auto
    83 
    84 lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
    85 by simp
    86 
    87 lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
    88 by blast
    89 
    90 lemma obj_sumE_f':
    91 "\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f x \<longrightarrow> P"
    92 by (cases x) blast+
    93 
    94 lemma obj_sumE_f:
    95 "\<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"
    96 by (rule allI) (rule obj_sumE_f')
    97 
    98 lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
    99 by (cases s) auto
   100 
   101 lemma obj_sum_step':
   102 "\<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"
   103 by (cases x) blast+
   104 
   105 lemma obj_sum_step:
   106 "\<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"
   107 by (rule allI) (rule obj_sum_step')
   108 
   109 lemma sum_case_if:
   110 "sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
   111 by simp
   112 
   113 ML_file "Tools/bnf_fp_util.ML"
   114 ML_file "Tools/bnf_fp_sugar_tactics.ML"
   115 ML_file "Tools/bnf_fp_sugar.ML"
   116 
   117 end