src/HOL/Lifting_Option.thy
author traytel
Tue, 13 Aug 2013 18:22:55 +0200
changeset 54163 e1a548c11845
parent 54149 cb82606b8215
child 56431 181751ad852f
permissions -rw-r--r--
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
     1 (*  Title:      HOL/Lifting_Option.thy
     2     Author:     Brian Huffman and Ondrej Kuncar
     3 *)
     4 
     5 header {* Setup for Lifting/Transfer for the option type *}
     6 
     7 theory Lifting_Option
     8 imports Lifting
     9 begin
    10 
    11 subsection {* Relator and predicator properties *}
    12 
    13 definition
    14   option_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool"
    15 where
    16   "option_rel R x y = (case (x, y) of (None, None) \<Rightarrow> True
    17     | (Some x, Some y) \<Rightarrow> R x y
    18     | _ \<Rightarrow> False)"
    19 
    20 lemma option_rel_simps[simp]:
    21   "option_rel R None None = True"
    22   "option_rel R (Some x) None = False"
    23   "option_rel R None (Some y) = False"
    24   "option_rel R (Some x) (Some y) = R x y"
    25   unfolding option_rel_def by simp_all
    26 
    27 abbreviation (input) option_pred :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool" where
    28   "option_pred \<equiv> option_case True"
    29 
    30 lemma option_rel_eq [relator_eq]:
    31   "option_rel (op =) = (op =)"
    32   by (simp add: option_rel_def fun_eq_iff split: option.split)
    33 
    34 lemma option_rel_mono[relator_mono]:
    35   assumes "A \<le> B"
    36   shows "(option_rel A) \<le> (option_rel B)"
    37 using assms by (auto simp: option_rel_def split: option.splits)
    38 
    39 lemma option_rel_OO[relator_distr]:
    40   "(option_rel A) OO (option_rel B) = option_rel (A OO B)"
    41 by (rule ext)+ (auto simp: option_rel_def OO_def split: option.split)
    42 
    43 lemma Domainp_option[relator_domain]:
    44   assumes "Domainp A = P"
    45   shows "Domainp (option_rel A) = (option_pred P)"
    46 using assms unfolding Domainp_iff[abs_def] option_rel_def[abs_def]
    47 by (auto iff: fun_eq_iff split: option.split)
    48 
    49 lemma reflp_option_rel[reflexivity_rule]:
    50   "reflp R \<Longrightarrow> reflp (option_rel R)"
    51   unfolding reflp_def split_option_all by simp
    52 
    53 lemma left_total_option_rel[reflexivity_rule]:
    54   "left_total R \<Longrightarrow> left_total (option_rel R)"
    55   unfolding left_total_def split_option_all split_option_ex by simp
    56 
    57 lemma left_unique_option_rel [reflexivity_rule]:
    58   "left_unique R \<Longrightarrow> left_unique (option_rel R)"
    59   unfolding left_unique_def split_option_all by simp
    60 
    61 lemma right_total_option_rel [transfer_rule]:
    62   "right_total R \<Longrightarrow> right_total (option_rel R)"
    63   unfolding right_total_def split_option_all split_option_ex by simp
    64 
    65 lemma right_unique_option_rel [transfer_rule]:
    66   "right_unique R \<Longrightarrow> right_unique (option_rel R)"
    67   unfolding right_unique_def split_option_all by simp
    68 
    69 lemma bi_total_option_rel [transfer_rule]:
    70   "bi_total R \<Longrightarrow> bi_total (option_rel R)"
    71   unfolding bi_total_def split_option_all split_option_ex by simp
    72 
    73 lemma bi_unique_option_rel [transfer_rule]:
    74   "bi_unique R \<Longrightarrow> bi_unique (option_rel R)"
    75   unfolding bi_unique_def split_option_all by simp
    76 
    77 lemma option_invariant_commute [invariant_commute]:
    78   "option_rel (Lifting.invariant P) = Lifting.invariant (option_pred P)"
    79   by (auto simp add: fun_eq_iff Lifting.invariant_def split_option_all)
    80 
    81 subsection {* Quotient theorem for the Lifting package *}
    82 
    83 lemma Quotient_option[quot_map]:
    84   assumes "Quotient R Abs Rep T"
    85   shows "Quotient (option_rel R) (Option.map Abs)
    86     (Option.map Rep) (option_rel T)"
    87   using assms unfolding Quotient_alt_def option_rel_def
    88   by (simp split: option.split)
    89 
    90 subsection {* Transfer rules for the Transfer package *}
    91 
    92 context
    93 begin
    94 interpretation lifting_syntax .
    95 
    96 lemma None_transfer [transfer_rule]: "(option_rel A) None None"
    97   by simp
    98 
    99 lemma Some_transfer [transfer_rule]: "(A ===> option_rel A) Some Some"
   100   unfolding fun_rel_def by simp
   101 
   102 lemma option_case_transfer [transfer_rule]:
   103   "(B ===> (A ===> B) ===> option_rel A ===> B) option_case option_case"
   104   unfolding fun_rel_def split_option_all by simp
   105 
   106 lemma option_map_transfer [transfer_rule]:
   107   "((A ===> B) ===> option_rel A ===> option_rel B) Option.map Option.map"
   108   unfolding Option.map_def by transfer_prover
   109 
   110 lemma option_bind_transfer [transfer_rule]:
   111   "(option_rel A ===> (A ===> option_rel B) ===> option_rel B)
   112     Option.bind Option.bind"
   113   unfolding fun_rel_def split_option_all by simp
   114 
   115 end
   116 
   117 end
   118