src/HOL/Transfer.thy
author Andreas Lochbihler
Thu, 26 Sep 2013 15:50:33 +0200
changeset 53927 abe2b313f0e5
parent 53011 aeee0a4be6cf
child 53944 50c8f7f21327
permissions -rw-r--r--
add lemmas

(*  Title:      HOL/Transfer.thy
    Author:     Brian Huffman, TU Muenchen
    Author:     Ondrej Kuncar, TU Muenchen
*)

header {* Generic theorem transfer using relations *}

theory Transfer
imports Hilbert_Choice
begin

subsection {* Relator for function space *}

definition
  fun_rel :: "('a \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> bool"
where
  "fun_rel A B = (\<lambda>f g. \<forall>x y. A x y \<longrightarrow> B (f x) (g y))"

locale lifting_syntax
begin
  notation fun_rel (infixr "===>" 55)
  notation map_fun (infixr "--->" 55)
end

context
begin
interpretation lifting_syntax .

lemma fun_relI [intro]:
  assumes "\<And>x y. A x y \<Longrightarrow> B (f x) (g y)"
  shows "(A ===> B) f g"
  using assms by (simp add: fun_rel_def)

lemma fun_relD:
  assumes "(A ===> B) f g" and "A x y"
  shows "B (f x) (g y)"
  using assms by (simp add: fun_rel_def)

lemma fun_relD2:
  assumes "(A ===> B) f g" and "A x x"
  shows "B (f x) (g x)"
  using assms unfolding fun_rel_def by auto

lemma fun_relE:
  assumes "(A ===> B) f g" and "A x y"
  obtains "B (f x) (g y)"
  using assms by (simp add: fun_rel_def)

lemma fun_rel_eq:
  shows "((op =) ===> (op =)) = (op =)"
  by (auto simp add: fun_eq_iff elim: fun_relE)

lemma fun_rel_eq_rel:
  shows "((op =) ===> R) = (\<lambda>f g. \<forall>x. R (f x) (g x))"
  by (simp add: fun_rel_def)


subsection {* Transfer method *}

text {* Explicit tag for relation membership allows for
  backward proof methods. *}

definition Rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
  where "Rel r \<equiv> r"

text {* Handling of equality relations *}

definition is_equality :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
  where "is_equality R \<longleftrightarrow> R = (op =)"

lemma is_equality_eq: "is_equality (op =)"
  unfolding is_equality_def by simp

text {* Reverse implication for monotonicity rules *}

definition rev_implies where
  "rev_implies x y \<longleftrightarrow> (y \<longrightarrow> x)"

text {* Handling of meta-logic connectives *}

definition transfer_forall where
  "transfer_forall \<equiv> All"

definition transfer_implies where
  "transfer_implies \<equiv> op \<longrightarrow>"

definition transfer_bforall :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
  where "transfer_bforall \<equiv> (\<lambda>P Q. \<forall>x. P x \<longrightarrow> Q x)"

lemma transfer_forall_eq: "(\<And>x. P x) \<equiv> Trueprop (transfer_forall (\<lambda>x. P x))"
  unfolding atomize_all transfer_forall_def ..

lemma transfer_implies_eq: "(A \<Longrightarrow> B) \<equiv> Trueprop (transfer_implies A B)"
  unfolding atomize_imp transfer_implies_def ..

lemma transfer_bforall_unfold:
  "Trueprop (transfer_bforall P (\<lambda>x. Q x)) \<equiv> (\<And>x. P x \<Longrightarrow> Q x)"
  unfolding transfer_bforall_def atomize_imp atomize_all ..

lemma transfer_start: "\<lbrakk>P; Rel (op =) P Q\<rbrakk> \<Longrightarrow> Q"
  unfolding Rel_def by simp

lemma transfer_start': "\<lbrakk>P; Rel (op \<longrightarrow>) P Q\<rbrakk> \<Longrightarrow> Q"
  unfolding Rel_def by simp

lemma transfer_prover_start: "\<lbrakk>x = x'; Rel R x' y\<rbrakk> \<Longrightarrow> Rel R x y"
  by simp

lemma untransfer_start: "\<lbrakk>Q; Rel (op =) P Q\<rbrakk> \<Longrightarrow> P"
  unfolding Rel_def by simp

lemma Rel_eq_refl: "Rel (op =) x x"
  unfolding Rel_def ..

lemma Rel_app:
  assumes "Rel (A ===> B) f g" and "Rel A x y"
  shows "Rel B (f x) (g y)"
  using assms unfolding Rel_def fun_rel_def by fast

lemma Rel_abs:
  assumes "\<And>x y. Rel A x y \<Longrightarrow> Rel B (f x) (g y)"
  shows "Rel (A ===> B) (\<lambda>x. f x) (\<lambda>y. g y)"
  using assms unfolding Rel_def fun_rel_def by fast

end

ML_file "Tools/transfer.ML"
setup Transfer.setup

declare refl [transfer_rule]

declare fun_rel_eq [relator_eq]

hide_const (open) Rel

context
begin
interpretation lifting_syntax .

text {* Handling of domains *}

lemma Domaimp_refl[transfer_domain_rule]:
  "Domainp T = Domainp T" ..

subsection {* Predicates on relations, i.e. ``class constraints'' *}

definition right_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
  where "right_total R \<longleftrightarrow> (\<forall>y. \<exists>x. R x y)"

definition right_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
  where "right_unique R \<longleftrightarrow> (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z)"

definition bi_total :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
  where "bi_total R \<longleftrightarrow> (\<forall>x. \<exists>y. R x y) \<and> (\<forall>y. \<exists>x. R x y)"

definition bi_unique :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
  where "bi_unique R \<longleftrightarrow>
    (\<forall>x y z. R x y \<longrightarrow> R x z \<longrightarrow> y = z) \<and>
    (\<forall>x y z. R x z \<longrightarrow> R y z \<longrightarrow> x = y)"

lemma bi_uniqueDr: "\<lbrakk> bi_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"
by(simp add: bi_unique_def)

lemma bi_uniqueDl: "\<lbrakk> bi_unique A; A x y; A z y \<rbrakk> \<Longrightarrow> x = z"
by(simp add: bi_unique_def)

lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A"
unfolding right_unique_def by blast

lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"
unfolding right_unique_def by blast

lemma right_total_alt_def:
  "right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All"
  unfolding right_total_def fun_rel_def
  apply (rule iffI, fast)
  apply (rule allI)
  apply (drule_tac x="\<lambda>x. True" in spec)
  apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)
  apply fast
  done

lemma right_unique_alt_def:
  "right_unique R \<longleftrightarrow> (R ===> R ===> op \<longrightarrow>) (op =) (op =)"
  unfolding right_unique_def fun_rel_def by auto

lemma bi_total_alt_def:
  "bi_total R \<longleftrightarrow> ((R ===> op =) ===> op =) All All"
  unfolding bi_total_def fun_rel_def
  apply (rule iffI, fast)
  apply safe
  apply (drule_tac x="\<lambda>x. \<exists>y. R x y" in spec)
  apply (drule_tac x="\<lambda>y. True" in spec)
  apply fast
  apply (drule_tac x="\<lambda>x. True" in spec)
  apply (drule_tac x="\<lambda>y. \<exists>x. R x y" in spec)
  apply fast
  done

lemma bi_unique_alt_def:
  "bi_unique R \<longleftrightarrow> (R ===> R ===> op =) (op =) (op =)"
  unfolding bi_unique_def fun_rel_def by auto

text {* Properties are preserved by relation composition. *}

lemma OO_def: "R OO S = (\<lambda>x z. \<exists>y. R x y \<and> S y z)"
  by auto

lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)"
  unfolding bi_total_def OO_def by metis

lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)"
  unfolding bi_unique_def OO_def by metis

lemma right_total_OO:
  "\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)"
  unfolding right_total_def OO_def by metis

lemma right_unique_OO:
  "\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)"
  unfolding right_unique_def OO_def by metis


subsection {* Properties of relators *}

lemma right_total_eq [transfer_rule]: "right_total (op =)"
  unfolding right_total_def by simp

lemma right_unique_eq [transfer_rule]: "right_unique (op =)"
  unfolding right_unique_def by simp

lemma bi_total_eq [transfer_rule]: "bi_total (op =)"
  unfolding bi_total_def by simp

lemma bi_unique_eq [transfer_rule]: "bi_unique (op =)"
  unfolding bi_unique_def by simp

lemma right_total_fun [transfer_rule]:
  "\<lbrakk>right_unique A; right_total B\<rbrakk> \<Longrightarrow> right_total (A ===> B)"
  unfolding right_total_def fun_rel_def
  apply (rule allI, rename_tac g)
  apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)
  apply clarify
  apply (subgoal_tac "(THE y. A x y) = y", simp)
  apply (rule someI_ex)
  apply (simp)
  apply (rule the_equality)
  apply assumption
  apply (simp add: right_unique_def)
  done

lemma right_unique_fun [transfer_rule]:
  "\<lbrakk>right_total A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A ===> B)"
  unfolding right_total_def right_unique_def fun_rel_def
  by (clarify, rule ext, fast)

lemma bi_total_fun [transfer_rule]:
  "\<lbrakk>bi_unique A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A ===> B)"
  unfolding bi_total_def fun_rel_def
  apply safe
  apply (rename_tac f)
  apply (rule_tac x="\<lambda>y. SOME z. B (f (THE x. A x y)) z" in exI)
  apply clarify
  apply (subgoal_tac "(THE x. A x y) = x", simp)
  apply (rule someI_ex)
  apply (simp)
  apply (rule the_equality)
  apply assumption
  apply (simp add: bi_unique_def)
  apply (rename_tac g)
  apply (rule_tac x="\<lambda>x. SOME z. B z (g (THE y. A x y))" in exI)
  apply clarify
  apply (subgoal_tac "(THE y. A x y) = y", simp)
  apply (rule someI_ex)
  apply (simp)
  apply (rule the_equality)
  apply assumption
  apply (simp add: bi_unique_def)
  done

lemma bi_unique_fun [transfer_rule]:
  "\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)"
  unfolding bi_total_def bi_unique_def fun_rel_def fun_eq_iff
  by (safe, metis, fast)


subsection {* Transfer rules *}

text {* Transfer rules using implication instead of equality on booleans. *}

lemma transfer_forall_transfer [transfer_rule]:
  "bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall"
  "right_total A \<Longrightarrow> ((A ===> op =) ===> implies) transfer_forall transfer_forall"
  "right_total A \<Longrightarrow> ((A ===> implies) ===> implies) transfer_forall transfer_forall"
  "bi_total A \<Longrightarrow> ((A ===> op =) ===> rev_implies) transfer_forall transfer_forall"
  "bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall"
  unfolding transfer_forall_def rev_implies_def fun_rel_def right_total_def bi_total_def
  by metis+

lemma transfer_implies_transfer [transfer_rule]:
  "(op =        ===> op =        ===> op =       ) transfer_implies transfer_implies"
  "(rev_implies ===> implies     ===> implies    ) transfer_implies transfer_implies"
  "(rev_implies ===> op =        ===> implies    ) transfer_implies transfer_implies"
  "(op =        ===> implies     ===> implies    ) transfer_implies transfer_implies"
  "(op =        ===> op =        ===> implies    ) transfer_implies transfer_implies"
  "(implies     ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"
  "(implies     ===> op =        ===> rev_implies) transfer_implies transfer_implies"
  "(op =        ===> rev_implies ===> rev_implies) transfer_implies transfer_implies"
  "(op =        ===> op =        ===> rev_implies) transfer_implies transfer_implies"
  unfolding transfer_implies_def rev_implies_def fun_rel_def by auto

lemma eq_imp_transfer [transfer_rule]:
  "right_unique A \<Longrightarrow> (A ===> A ===> op \<longrightarrow>) (op =) (op =)"
  unfolding right_unique_alt_def .

lemma eq_transfer [transfer_rule]:
  assumes "bi_unique A"
  shows "(A ===> A ===> op =) (op =) (op =)"
  using assms unfolding bi_unique_def fun_rel_def by auto

lemma Domainp_iff: "Domainp T x \<longleftrightarrow> (\<exists>y. T x y)"
  by auto

lemma right_total_Ex_transfer[transfer_rule]:
  assumes "right_total A"
  shows "((A ===> op=) ===> op=) (Bex (Collect (Domainp A))) Ex"
using assms unfolding right_total_def Bex_def fun_rel_def Domainp_iff[abs_def]
by blast

lemma right_total_All_transfer[transfer_rule]:
  assumes "right_total A"
  shows "((A ===> op =) ===> op =) (Ball (Collect (Domainp A))) All"
using assms unfolding right_total_def Ball_def fun_rel_def Domainp_iff[abs_def]
by blast

lemma All_transfer [transfer_rule]:
  assumes "bi_total A"
  shows "((A ===> op =) ===> op =) All All"
  using assms unfolding bi_total_def fun_rel_def by fast

lemma Ex_transfer [transfer_rule]:
  assumes "bi_total A"
  shows "((A ===> op =) ===> op =) Ex Ex"
  using assms unfolding bi_total_def fun_rel_def by fast

lemma If_transfer [transfer_rule]: "(op = ===> A ===> A ===> A) If If"
  unfolding fun_rel_def by simp

lemma Let_transfer [transfer_rule]: "(A ===> (A ===> B) ===> B) Let Let"
  unfolding fun_rel_def by simp

lemma id_transfer [transfer_rule]: "(A ===> A) id id"
  unfolding fun_rel_def by simp

lemma comp_transfer [transfer_rule]:
  "((B ===> C) ===> (A ===> B) ===> (A ===> C)) (op \<circ>) (op \<circ>)"
  unfolding fun_rel_def by simp

lemma fun_upd_transfer [transfer_rule]:
  assumes [transfer_rule]: "bi_unique A"
  shows "((A ===> B) ===> A ===> B ===> A ===> B) fun_upd fun_upd"
  unfolding fun_upd_def [abs_def] by transfer_prover

lemma nat_case_transfer [transfer_rule]:
  "(A ===> (op = ===> A) ===> op = ===> A) nat_case nat_case"
  unfolding fun_rel_def by (simp split: nat.split)

lemma nat_rec_transfer [transfer_rule]:
  "(A ===> (op = ===> A ===> A) ===> op = ===> A) nat_rec nat_rec"
  unfolding fun_rel_def by (clarsimp, rename_tac n, induct_tac n, simp_all)

lemma funpow_transfer [transfer_rule]:
  "(op = ===> (A ===> A) ===> (A ===> A)) compow compow"
  unfolding funpow_def by transfer_prover

lemma Domainp_forall_transfer [transfer_rule]:
  assumes "right_total A"
  shows "((A ===> op =) ===> op =)
    (transfer_bforall (Domainp A)) transfer_forall"
  using assms unfolding right_total_def
  unfolding transfer_forall_def transfer_bforall_def fun_rel_def Domainp_iff
  by metis

lemma forall_transfer [transfer_rule]:
  "bi_total A \<Longrightarrow> ((A ===> op =) ===> op =) transfer_forall transfer_forall"
  unfolding transfer_forall_def by (rule All_transfer)

end

end