Theory Bounded_Quantifiers

✐‹creator "Kevin Kappelmann"›
subsection ‹Bounded Quantifiers›
theory Bounded_Quantifiers
  imports
    HOL_Basics_Base
    Predicates_Lattice
    ML_Unification.ML_Unification_HOL_Setup
begin

consts ball :: "'a  ('b  bool)  bool"

bundle ball_syntax
begin
syntax
  "_ball"  :: [idts, 'a, bool]  bool ("(2_ : _./ _)" 10)
  "_ball2" :: [idts, 'a, bool]  bool
notation ball ("(_)")
end
bundle no_ball_syntax
begin
no_syntax
  "_ball"  :: [idts, 'a, bool]  bool ("(2_ : _./ _)" 10)
  "_ball2" :: [idts, 'a, bool]  bool
no_notation ball ("(_)")
end
unbundle ball_syntax
translations
  "x xs : P. Q"  "CONST ball P (λx. _ball2 xs P Q)"
  "_ball2 x P Q"  "x : P. Q"
  "x : P. Q"  "CONST ball P (λx. Q)"

consts bex :: "'a  ('b  bool)  bool"

bundle bex_syntax
begin
syntax
  "_bex"  :: [idts, 'a, bool]  bool ("(2_ : _./ _)" 10)
  "_bex2" :: [idts, 'a, bool]  bool
notation bex ("(_)")
end
bundle no_bex_syntax
begin
no_syntax
  "_bex"  :: [idts, 'a, bool]  bool ("(2_ : _./ _)" 10)
  "_bex2" :: [idts, 'a, bool]  bool
no_notation bex ("(_)")
end
unbundle bex_syntax
translations
  "x xs : P. Q"  "CONST bex P (λx. _bex2 xs P Q)"
  "_bex2 x P Q"  "x : P. Q"
  "x : P. Q"  "CONST bex P (λx. Q)"

consts bex1 :: "'a  ('b  bool)  bool"

bundle bex1_syntax
begin
syntax
  "_bex1"  :: [idts, 'a, bool]  bool ("(2∃!_ : _./ _)" 10)
  "_bex12" :: [idts, 'a, bool]  bool
notation bex1 ("∃!(_)")
end
bundle no_bex1_syntax
begin
no_syntax
  "_bex1"  :: [idts, 'a, bool]  bool ("(2∃!_ : _./ _)" 10)
  "_bex12" :: [idts, 'a, bool]  bool
no_notation bex1 ("∃!(_)")
end
unbundle bex1_syntax
translations
  "∃!x xs : P. Q"  "CONST bex1 P (λx. _bex12 xs P Q)"
  "_bex12 x P Q"  "∃!x : P. Q"
  "∃!x : P. Q"  "CONST bex1 P (λx. Q)"

bundle bounded_quantifier_syntax begin unbundle ball_syntax bex_syntax bex1_syntax end
bundle no_bounded_quantifier_syntax begin unbundle no_ball_syntax no_bex_syntax no_bex1_syntax end

definition "ball_pred P Q  x. P x  Q x"
adhoc_overloading ball ball_pred

definition "bex_pred P Q  x. P x  Q x"
adhoc_overloading bex bex_pred

definition "bex1_pred P Q  ∃!x. P x  Q x"
adhoc_overloading bex1 bex1_pred

(*copied from HOL.Set.thy*)
simproc_setup defined_ball ("x : P. Q x  U x") = K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms ball_pred_def}))
simproc_setup defined_bex ("x : P. Q x  U x") = K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms bex_pred_def}))

lemma ballI [intro!]:
  assumes "x. P x  Q x"
  shows "x : P. Q x"
  using assms unfolding ball_pred_def by blast

lemma ballE [elim]:
  assumes "x : P. Q x"
  obtains "x. P x  Q x"
  using assms unfolding ball_pred_def by blast

lemma ballE':
  assumes "x : P. Q x"
  obtains "¬(P x)" | "P x" "Q x"
  using assms by blast

lemma ballD: "x : P. Q x  P x  Q x"
  by blast

lemma ball_cong: "P = P'; x. P' x  Q x  Q' x  (x : P. Q x)  (x : P'. Q' x)"
  by auto

lemma ball_cong_simp [cong]:
  "P = P'; x. P' x =simp=> Q x  Q' x  (x : P. Q x)  (x : P'. Q' x)"
  unfolding simp_implies_def by (rule ball_cong)

(*copied from HOL.Set.thy*)
ML structure Simpdata =
struct
  open Simpdata
  val mksimps_pairs = [(const_nameball_pred, @{thms ballD})] @ mksimps_pairs
end
open Simpdata
declaration fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))

lemma atomize_ball: "(x. P x  Q x)  Trueprop (x : P. Q x)"
  by (simp only: ball_pred_def atomize_all atomize_imp)

declare atomize_ball[symmetric, rulify]
declare atomize_ball[symmetric, defn]

lemma bexI [intro!]:
  (*better argument order: Q often determines the choice for x*)
  assumes "x. Q x  P x"
  shows "x : P. Q x"
  using assms unfolding bex_pred_def by blast

lemma bexE [elim!]:
  assumes "x : P. Q x"
  obtains x where "P x" "Q x"
  using assms unfolding bex_pred_def by blast

lemma bexD:
  assumes "x : P. Q x"
  shows "x. P x  Q x"
  using assms by blast

lemma bex_cong:
  "P = P'; x. P' x  Q x  Q' x  (x : P. Q x)  (x : P'. Q' x)"
  by blast

lemma bex_cong_simp [cong]:
  "P = P'; x. P' x =simp=> Q x  Q' x  (x : P. Q x)  (x : P'. Q' x)"
  unfolding simp_implies_def by (rule bex_cong)

lemma bex1I [intro]:
  (*better argument order: Q often determines the choice for x*)
  assumes "∃!x. Q x  P x"
  shows "∃!x : P. Q x"
  using assms unfolding bex1_pred_def by blast

lemma bex1D [dest!]:
  assumes "∃!x : P. Q x"
  shows "∃!x. P x  Q x"
  using assms unfolding bex1_pred_def by blast

lemma bex1_cong: "P = P'; x. P x  Q x  Q' x  (∃!x : P. Q x)  (∃!x : P'. Q' x)"
  by blast

lemma bex1_cong_simp [cong]:
  "P = P'; x. P x =simp=> Q x  Q' x  (∃!x : P. Q x)  (∃!x : P'. Q' x)"
  unfolding simp_implies_def by (rule bex1_cong)

lemma ball_iff_ex_pred [iff]: "(x : P. Q)  ((x. P x)  Q)"
  by auto

lemma bex_iff_ex_and [iff]: "(x : P. Q)  ((x. P x)  Q)"
  by blast

lemma ball_eq_imp_iff_imp [iff]: "(x : P. x = y  Q x)  (P y  Q y)"
  by blast

lemma ball_eq_imp_iff_imp' [iff]: "(x : P. y = x  Q x)  (P y  Q y)"
  by blast

lemma bex_eq_iff_pred [iff]: "(x : P. x = y)  P y"
  by blast

lemma bex_eq_iff_pred' [iff]: "(x : P. y = x)  P y"
  by blast

lemma bex_eq_and_iff_pred [iff]: "(x : P. x = y  Q x)  P y  Q y"
  by blast

lemma bex_eq_and_iff_pred' [iff]: "(x : P. y = x  Q x)  P y  Q y"
  by blast

lemma ball_and_iff_ball_and_ball: "(x : P. Q x  U x)  (x : P. Q x)  (x : P. U x)"
  by auto

lemma bex_or_iff_bex_or_bex: "(x : P. Q x  U x)  (x : P. Q x)  (x : P. U x)"
  by auto

lemma ball_or_iff_ball_or [iff]: "(x : P. Q x  U)  ((x : P. Q x)  U)"
  by auto

lemma ball_or_iff_or_ball [iff]: "(x : P. Q  U x)  (Q  (x : P. U x))"
  by auto

lemma ball_imp_iff_imp_ball [iff]: "(x : P. Q  U x)  (Q  (x : P. U x))"
  by auto

lemma bex_and_iff_bex_and [iff]: "(x : P. Q x  U)  ((x : P. Q x)  U)"
  by auto

lemma bex_and_iff_and_bex [iff]: "(x : P. Q  U x)  (Q  (x : P. U x))"
  by auto

lemma ball_imp_iff_bex_imp [iff]: "(x : P. Q x  U)  ((x : P. Q x)  U)"
  by auto

lemma not_ball_iff_bex_not [iff]: "(¬(x : P. Q x))  (x : P. ¬(Q x))"
  by auto

lemma not_bex_iff_ball_not [iff]: "(¬(x : P. Q x))  (x : P. ¬(Q x))"
  by auto

lemma bex1_iff_bex_and [iff]: "(∃!x : P. Q)  ((∃!x. P x)  Q)"
  by auto

lemma ball_bottom_eq_top [simp]: "= " by auto
lemma bex_bottom_eq_bottom [simp]: "= " by fastforce
lemma bex1_bottom_eq_bottom [simp]: "∃!= " by fastforce

lemma ball_top_eq_all [simp]: "= All" by fastforce

lemma ball_top_eq_all_uhint [uhint]:
  assumes "P  ( :: 'a  bool)"
  shows "P All"
  using assms by simp

lemma bex_top_eq_ex [simp]: "= Ex" by fastforce

lemma bex_top_eq_ex_uhint [uhint]:
  assumes "P  ( :: 'a  bool)"
  shows "P Ex"
  using assms by simp

lemma bex1_top_eq_ex1 [simp]: "∃!= Ex1" by fastforce

lemma bex1_top_eq_ex1_uhint [uhint]:
  assumes "P  ( :: 'a  bool)"
  shows "∃!P Ex1"
  using assms by simp

end