Theory Powerset_Axiom

section‹The Powerset Axiom in $M[G]$›
theory Powerset_Axiom
  imports Renaming_Auto Separation_Axiom Pairing_Axiom Union_Axiom
begin

simple_rename "perm_pow" src "[ss,p,l,o,fs,χ]" tgt "[fs,ss,sp,p,l,o,χ]"

lemma Collect_inter_Transset:
  assumes
    "Transset(M)" "b  M"
  shows
    "{xb . P(x)} = {xb . P(x)}  M"
  using assms unfolding Transset_def
  by (auto)

context G_generic  begin

lemma name_components_in_M:
  assumes "<σ,p>θ" "θ  M"
  shows   "σM" "pM"
proof -
  from assms obtain a where
    "σ  a" "p  a" "a<σ,p>"
    unfolding Pair_def by auto
  moreover from assms
  have "<σ,p>M"
    using transitivity by simp
  moreover from calculation
  have "aM"
    using transitivity by simp
  ultimately
  show "σM" "pM"
    using transitivity by simp_all
qed

lemma sats_fst_snd_in_M:
  assumes
    "AM" "BM" "φ  formula" "pM" "lM" "oM" "χM"
    "arity(φ)  6"
  shows
    "{sq A×B . sats(M,φ,[snd(sq),p,l,o,fst(sq),χ])}  M"
    (is "  M")
proof -
  have "6nat" "7nat" by simp_all
  let ?φ' = "ren(φ)`6`7`perm_pow_fn"
  from AM BM have
    "A×B  M"
    using cartprod_closed by simp
  from arity(φ)  6 φ formula 6_ 7_
  have "?φ'  formula" "arity(?φ')7"
    unfolding perm_pow_fn_def
    using  perm_pow_thm  arity_ren ren_tc Nil_type
    by auto
  with ?φ'  formula
  have 1: "arity(Exists(Exists(And(pair_fm(0,1,2),?φ'))))5"     (is "arity()5")
    unfolding pair_fm_def upair_fm_def
    using nat_simp_union pred_le arity_type by auto
  {
    fix sp
    note A×B  M
    moreover
    assume "sp  A×B"
    moreover from calculation
    have "fst(sp)  A" "snd(sp)  B"
      using fst_type snd_type by simp_all
    ultimately
    have "sp  M" "fst(sp)  M" "snd(sp)  M"
      using  AM BM transitivity
      by simp_all
    note inM = AM BM pM lM oM χM
      spM fst(sp)M snd(sp)M
    with 1 sp  M ?φ'  formula
    have "M, [sp,p,l,o,χ]@[p]    M,[sp,p,l,o,χ]  " (is "M,?env0@ __  _")
      using arity_sats_iff[of  "[p]" M ?env0] by auto
    also from inM sp  A×B
    have "...  sats(M,?φ',[fst(sp),snd(sp),sp,p,l,o,χ])"
      by auto
    also from inM φ  formula arity(φ)  6
    have "...  sats(M,φ,[snd(sp),p,l,o,fst(sp),χ])"
      (is "sats(_,_,?env1)  sats(_,_,?env2)")
      using sats_iff_sats_ren[of φ 6 7 ?env2 M ?env1 perm_pow_fn] perm_pow_thm
      unfolding perm_pow_fn_def by simp
    finally
    have "sats(M,,[sp,p,l,o,χ,p])  sats(M,φ,[snd(sp),p,l,o,fst(sp),χ])"
      by simp
  }
  then have
    " = {spA×B . sats(M,,[sp,p,l,o,χ,p])}"
    by auto
  also from assms A×BM have
    " ...  M"
  proof -
    from 1
    have "arity()  6"
      using leI by simp
    moreover from ?φ'  formula
    have "  formula"
      by simp
    moreover note assms A×BM
    ultimately 
    show "{x  A×B . sats(M, , [x, p, l, o, χ, p])}  M"
      using separation_ax separation_iff
      by simp
  qed
  finally show ?thesis .
qed

lemma Pow_inter_MG:
  assumes
    "aM[G]"
  shows
    "Pow(a)  M[G]  M[G]"
proof -
  from assms obtain τ where
    "τ  M" "val(G, τ) = a"
    using GenExtD by auto
  let ?Q="Pow(domain(τ)×P)  M"
  from τM 
  have "domain(τ)×P  M" "domain(τ)  M"
    using domain_closed cartprod_closed P_in_M
    by simp_all
  then 
  have "?Q  M"
  proof -
    from power_ax domain(τ)×P  M obtain Q where
      "powerset(##M,domain(τ)×P,Q)" "Q  M"
      unfolding power_ax_def by auto
    moreover from calculation 
    have "zQ  zM" for z
      using transitivity by blast
    ultimately
    have "Q = {aPow(domain(τ)×P) . aM}"
      using domain(τ)×P  M powerset_abs[of "domain(τ)×P" Q]
      by (simp flip: setclass_iff)
    also 
    have " ... = ?Q"
      by auto
    finally 
    show ?thesis using QM by simp
  qed
  let
    ="?Q×{one}"
  let
    ?b="val(G,)"
  from ?QM 
  have "M"
    using one_in_P P_in_M transitivity
    by (simp flip: setclass_iff)
  from M 
  have "?b  M[G]"
    using GenExtI by simp
  have "Pow(a)  M[G]  ?b"
  proof
    fix c
    assume "c  Pow(a)  M[G]"
    then obtain χ where
      "cM[G]" "χ  M" "val(G,χ) = c"
      using GenExtD by auto
    let ="{sp domain(τ)×P . snd(sp)  (Member(0,1)) [fst(sp),χ] }"
    have "arity(forces(Member(0,1))) = 6"
      using arity_forces_at by auto
    with domain(τ)  M χ  M 
    have "  M"
      using P_in_M one_in_M leq_in_M sats_fst_snd_in_M
      by simp
    then 
    have "  ?Q"
      by auto
    then 
    have "val(G,)  ?b"
      using one_in_G one_in_P generic val_of_elem [of  one  G]
      by auto
    have "val(G,) = c"
    proof(intro equalityI subsetI)
      fix x
      assume "x  val(G,)"
      then obtain σ p where
        1: "<σ,p>" "pG" "val(G,σ) =  x"
        using elem_of_val_pair
        by blast
      moreover from <σ,p>   M
      have "σM"
        using name_components_in_M[of _ _ ] by auto
      moreover from 1 
      have "(p  (Member(0,1)) [σ,χ])" "pP"
        by simp_all
      moreover 
      note val(G,χ) = c
      ultimately 
      have "sats(M[G],Member(0,1),[x,c])"
        using χ  M generic definition_of_forcing nat_simp_union
        by auto
      moreover 
      have "xM[G]"
        using val(G,σ) =  x σM  χM GenExtI by blast
      ultimately 
      show "xc"
        using cM[G] by simp
    next
      fix x
      assume "x  c"
      with c  Pow(a)  M[G] 
      have "x  a" "cM[G]" "xM[G]"
        using transitivity_MG
        by auto
      with val(G, τ) = a 
      obtain σ where
        "σdomain(τ)" "val(G,σ) =  x"
        using elem_of_val
        by blast
      moreover note xc val(G,χ) = c
      moreover from calculation 
      have "val(G,σ)  val(G,χ)"
        by simp
      moreover note cM[G] xM[G]
      moreover from calculation 
      have "sats(M[G],Member(0,1),[x,c])"
        by simp
      moreover 
      have "Member(0,1)formula" by simp
      moreover 
      have "σM"
      proof -
        from σdomain(τ) 
        obtain p where "<σ,p>  τ"
          by auto
        with τM 
        show ?thesis
          using name_components_in_M by blast
      qed
      moreover note χ  M
      ultimately 
      obtain p where "pG" "(p  Member(0,1) [σ,χ])"
        using generic truth_lemma[of "Member(0,1)" "G" "[σ,χ]" ] nat_simp_union
        by auto
      moreover from pG 
      have "pP"
        using generic unfolding M_generic_def filter_def by blast
      ultimately
      have "<σ,p>"
        using σdomain(τ) by simp
      with val(G,σ) =  x pG 
      show "xval(G,)"
        using val_of_elem [of _ _ ""] by auto
    qed
    with val(G,)  ?b 
    show "c?b" by simp
  qed
  then 
  have "Pow(a)  M[G] = {x?b . xa & xM[G]}"
    by auto
  also from aM[G] 
  have " ... = {x?b . sats(M[G],subset_fm(0,1),[x,a]) & xM[G]}"
    using Transset_MG by force
  also 
  have " ... = {x?b . sats(M[G],subset_fm(0,1),[x,a])}  M[G]"
    by auto
  also from ?bM[G] 
  have " ... = {x?b . sats(M[G],subset_fm(0,1),[x,a])}"
    using Collect_inter_Transset Transset_MG
    by simp
  also from ?bM[G] aM[G]
  have " ...  M[G]"
    using Collect_sats_in_MG GenExtI nat_simp_union by simp
  finally show ?thesis .
qed
end (* context: G_generic *)


context G_generic begin

interpretation mgtriv: M_trivial "##M[G]"
  using generic Union_MG pairing_in_MG zero_in_MG transitivity_MG
  unfolding M_trivial_def M_trans_def M_trivial_axioms_def by (simp; blast)


theorem power_in_MG : "power_ax(##(M[G]))"
  unfolding power_ax_def
proof (intro rallI, simp only:setclass_iff rex_setclass_is_bex)
  (* After simplification, we have to show that for every
     a∈M[G] there exists some x∈M[G] with powerset(##M[G],a,x)
  *)
  fix a
  assume "a  M[G]"
  then
  have "(##M[G])(a)" by simp
  have "{xPow(a) . x  M[G]} = Pow(a)  M[G]"
    by auto
  also from aM[G] 
  have " ...  M[G]"
    using Pow_inter_MG by simp
  finally 
  have "{xPow(a) . x  M[G]}  M[G]" .
  moreover from aM[G] {xPow(a) . x  M[G]}  _ 
  have "powerset(##M[G], a, {xPow(a) . x  M[G]})"
    using mgtriv.powerset_abs[OF (##M[G])(a)]
    by simp
  ultimately 
  show "xM[G] . powerset(##M[G], a, x)"
    by auto
qed
end (* context: G_generic *)
end