src/HOL/Library/Quotient_Option.thy
changeset 53012 cb82606b8215
parent 53010 ec5e6f69bd65
child 53026 e1a548c11845
--- a/src/HOL/Library/Quotient_Option.thy	Tue Aug 13 15:59:22 2013 +0200
+++ b/src/HOL/Library/Quotient_Option.thy	Tue Aug 13 15:59:22 2013 +0200
@@ -1,5 +1,5 @@
 (*  Title:      HOL/Library/Quotient_Option.thy
-    Author:     Cezary Kaliszyk, Christian Urban and Brian Huffman
+    Author:     Cezary Kaliszyk and Christian Urban
 *)
 
 header {* Quotient infrastructure for the option type *}
@@ -8,31 +8,7 @@
 imports Main Quotient_Syntax
 begin
 
-subsection {* Relator for option type *}
-
-fun
-  option_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool"
-where
-  "option_rel R None None = True"
-| "option_rel R (Some x) None = False"
-| "option_rel R None (Some x) = False"
-| "option_rel R (Some x) (Some y) = R x y"
-
-lemma option_rel_unfold:
-  "option_rel R x y = (case (x, y) of (None, None) \<Rightarrow> True
-    | (Some x, Some y) \<Rightarrow> R x y
-    | _ \<Rightarrow> False)"
-  by (cases x) (cases y, simp_all)+
-
-fun option_pred :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool"
-where
-  "option_pred R None = True"
-| "option_pred R (Some x) = R x"
-
-lemma option_pred_unfold:
-  "option_pred P x = (case x of None \<Rightarrow> True
-    | Some x \<Rightarrow> P x)"
-by (cases x) simp_all
+subsection {* Rules for the Quotient package *}
 
 lemma option_rel_map1:
   "option_rel R (Option.map f x) y \<longleftrightarrow> option_rel (\<lambda>x. R (f x)) x y"
@@ -46,37 +22,10 @@
   "Option.map id = id"
   by (simp add: id_def Option.map.identity fun_eq_iff)
 
-lemma option_rel_eq [id_simps, relator_eq]:
+lemma option_rel_eq [id_simps]:
   "option_rel (op =) = (op =)"
   by (simp add: option_rel_unfold fun_eq_iff split: option.split)
 
-lemma option_rel_mono[relator_mono]:
-  assumes "A \<le> B"
-  shows "(option_rel A) \<le> (option_rel B)"
-using assms by (auto simp: option_rel_unfold split: option.splits)
-
-lemma option_rel_OO[relator_distr]:
-  "(option_rel A) OO (option_rel B) = option_rel (A OO B)"
-by (rule ext)+ (auto simp: option_rel_unfold OO_def split: option.split)
-
-lemma Domainp_option[relator_domain]:
-  assumes "Domainp A = P"
-  shows "Domainp (option_rel A) = (option_pred P)"
-using assms unfolding Domainp_iff[abs_def] option_rel_unfold[abs_def] option_pred_unfold[abs_def]
-by (auto iff: fun_eq_iff split: option.split)
-
-lemma reflp_option_rel[reflexivity_rule]:
-  "reflp R \<Longrightarrow> reflp (option_rel R)"
-  unfolding reflp_def split_option_all by simp
-
-lemma left_total_option_rel[reflexivity_rule]:
-  "left_total R \<Longrightarrow> left_total (option_rel R)"
-  unfolding left_total_def split_option_all split_option_ex by simp
-
-lemma left_unique_option_rel [reflexivity_rule]:
-  "left_unique R \<Longrightarrow> left_unique (option_rel R)"
-  unfolding left_unique_def split_option_all by simp
-
 lemma option_symp:
   "symp R \<Longrightarrow> symp (option_rel R)"
   unfolding symp_def split_option_all option_rel.simps by fast
@@ -89,65 +38,6 @@
   "equivp R \<Longrightarrow> equivp (option_rel R)"
   by (blast intro: equivpI reflp_option_rel option_symp option_transp elim: equivpE)
 
-lemma right_total_option_rel [transfer_rule]:
-  "right_total R \<Longrightarrow> right_total (option_rel R)"
-  unfolding right_total_def split_option_all split_option_ex by simp
-
-lemma right_unique_option_rel [transfer_rule]:
-  "right_unique R \<Longrightarrow> right_unique (option_rel R)"
-  unfolding right_unique_def split_option_all by simp
-
-lemma bi_total_option_rel [transfer_rule]:
-  "bi_total R \<Longrightarrow> bi_total (option_rel R)"
-  unfolding bi_total_def split_option_all split_option_ex by simp
-
-lemma bi_unique_option_rel [transfer_rule]:
-  "bi_unique R \<Longrightarrow> bi_unique (option_rel R)"
-  unfolding bi_unique_def split_option_all by simp
-
-subsection {* Transfer rules for transfer package *}
-
-lemma None_transfer [transfer_rule]: "(option_rel A) None None"
-  by simp
-
-lemma Some_transfer [transfer_rule]: "(A ===> option_rel A) Some Some"
-  unfolding fun_rel_def by simp
-
-lemma option_case_transfer [transfer_rule]:
-  "(B ===> (A ===> B) ===> option_rel A ===> B) option_case option_case"
-  unfolding fun_rel_def split_option_all by simp
-
-lemma option_map_transfer [transfer_rule]:
-  "((A ===> B) ===> option_rel A ===> option_rel B) Option.map Option.map"
-  unfolding Option.map_def by transfer_prover
-
-lemma option_bind_transfer [transfer_rule]:
-  "(option_rel A ===> (A ===> option_rel B) ===> option_rel B)
-    Option.bind Option.bind"
-  unfolding fun_rel_def split_option_all by simp
-
-subsection {* Setup for lifting package *}
-
-lemma Quotient_option[quot_map]:
-  assumes "Quotient R Abs Rep T"
-  shows "Quotient (option_rel R) (Option.map Abs)
-    (Option.map Rep) (option_rel T)"
-  using assms unfolding Quotient_alt_def option_rel_unfold
-  by (simp split: option.split)
-
-lemma option_invariant_commute [invariant_commute]:
-  "option_rel (Lifting.invariant P) = Lifting.invariant (option_pred P)"
-  apply (simp add: fun_eq_iff Lifting.invariant_def)
-  apply (intro allI) 
-  apply (case_tac x rule: option.exhaust)
-  apply (case_tac xa rule: option.exhaust)
-  apply auto[2]
-  apply (case_tac xa rule: option.exhaust)
-  apply auto
-done
-
-subsection {* Rules for quotient package *}
-
 lemma option_quotient [quot_thm]:
   assumes "Quotient3 R Abs Rep"
   shows "Quotient3 (option_rel R) (Option.map Abs) (Option.map Rep)"