Theory Forcing_Theorems

section‹The Forcing Theorems›

theory Forcing_Theorems
  imports
    Forces_Definition

begin

context forcing_data
begin

subsection‹The forcing relation in context›

abbreviation Forces :: "[i, i, i]  o"  ("_  _ _" [36,36,36] 60) where
  "p  φ env      M, ([p,P,leq,one] @ env)  forces(φ)"

lemma Collect_forces :
  assumes 
    fty: "φformula" and
    far: "arity(φ)length(env)" and
    envty: "envlist(M)"
  shows
    "{pP . p  φ env}  M"
proof -
  have "zP  zM" for z
    using P_in_M transitivity[of z P] by simp
  moreover
  have "separation(##M,λp. (p  φ env))"
        using separation_ax arity_forces far fty P_in_M leq_in_M one_in_M envty arity_forces_le
    by simp
  then 
  have "Collect(P,λp. (p  φ env))M"
    using separation_closed P_in_M by simp
  then show ?thesis by simp
qed

lemma forces_mem_iff_dense_below:  "pP  forces_mem(p,t1,t2)  dense_below(
    {qP. s. r. rP  s,r  t2  qr  forces_eq(q,t1,s)}
    ,p)"
  using def_forces_mem[of p t1 t2] by blast

subsection‹Kunen 2013, Lemma IV.2.37(a)›

lemma strengthening_eq:
  assumes "pP" "rP" "rp" "forces_eq(p,t1,t2)"
  shows "forces_eq(r,t1,t2)"
  using assms def_forces_eq[of _ t1 t2] leq_transD by blast
(* Long proof *)
(*
proof -
  {
    fix s q
    assume "q≼ r" "q∈P"
    with assms
    have "q≼p"
      using leq_preord unfolding preorder_on_def trans_on_def by blast
    moreover 
    note ‹q∈P› assms
    moreover
    assume "s∈domain(t1) ∪ domain(t2)" 
    ultimately
    have "forces_mem(q, s, t1) ⟷ forces_mem(q, s, t2)"
      using def_forces_eq[of p t1 t2] by simp
  }
  with ‹r∈P›
  show ?thesis using def_forces_eq[of r t1 t2] by blast
qed
*)

subsection‹Kunen 2013, Lemma IV.2.37(a)›
lemma strengthening_mem: 
  assumes "pP" "rP" "rp" "forces_mem(p,t1,t2)"
  shows "forces_mem(r,t1,t2)"
  using assms forces_mem_iff_dense_below dense_below_under by auto

subsection‹Kunen 2013, Lemma IV.2.37(b)›
lemma density_mem: 
  assumes "pP"
  shows "forces_mem(p,t1,t2)   dense_below({qP. forces_mem(q,t1,t2)},p)"
proof
  assume "forces_mem(p,t1,t2)"
  with assms
  show "dense_below({qP. forces_mem(q,t1,t2)},p)"
    using forces_mem_iff_dense_below strengthening_mem[of p] ideal_dense_below by auto
next
  assume "dense_below({q  P . forces_mem(q, t1, t2)}, p)"
  with assms
  have "dense_below({qP. 
    dense_below({q'P. s r. r  P  s,rt2  q'r  forces_eq(q',t1,s)},q)
    },p)"
    using forces_mem_iff_dense_below by simp
  with assms
  show "forces_mem(p,t1,t2)"
    using dense_below_dense_below forces_mem_iff_dense_below[of p t1 t2] by blast
qed

lemma aux_density_eq:
  assumes 
    "dense_below(
    {q'P. q. qP  qq'  forces_mem(q,s,t1)  forces_mem(q,s,t2)}
    ,p)"
    "forces_mem(q,s,t1)" "qP" "pP" "qp"
  shows
    "dense_below({rP. forces_mem(r,s,t2)},q)"
proof
  fix r
  assume "rP" "rq"
  moreover from this and pP qp qP
  have "rp"
    using leq_transD by simp
  moreover
  note forces_mem(q,s,t1) dense_below(_,p) qP
  ultimately
  obtain q1 where "q1r" "q1P" "forces_mem(q1,s,t2)"
    using strengthening_mem[of q _ s t1] leq_reflI leq_transD[of _ r q] by blast
  then
  show "d{r  P . forces_mem(r, s, t2)}. d  P  d r"
    by blast
qed

(* Kunen 2013, Lemma IV.2.37(b) *)
lemma density_eq:
  assumes "pP"
  shows "forces_eq(p,t1,t2)   dense_below({qP. forces_eq(q,t1,t2)},p)"
proof
  assume "forces_eq(p,t1,t2)"
  with pP
  show "dense_below({qP. forces_eq(q,t1,t2)},p)"
    using strengthening_eq ideal_dense_below by auto
next
  assume "dense_below({qP. forces_eq(q,t1,t2)},p)"
  {
    fix s q 
    let ?D1="{q'P. sdomain(t1)  domain(t2). q. q  P  qq' 
           forces_mem(q,s,t1)forces_mem(q,s,t2)}"
    let ?D2="{q'P. q. qP  qq'  forces_mem(q,s,t1)  forces_mem(q,s,t2)}"
    assume "sdomain(t1)  domain(t2)" 
    then
    have "?D1?D2" by blast
    with dense_below(_,p)
    have "dense_below({q'P. sdomain(t1)  domain(t2). q. q  P  qq' 
           forces_mem(q,s,t1)forces_mem(q,s,t2)},p)"
      using dense_below_cong'[OF pP def_forces_eq[of _ t1 t2]] by simp
    with pP ?D1?D2
    have "dense_below({q'P. q. qP  qq'  
            forces_mem(q,s,t1)  forces_mem(q,s,t2)},p)"
      using dense_below_mono by simp
    moreover from this (* Automatic tools can't handle this symmetry 
                          in order to apply aux_density_eq below *)
    have  "dense_below({q'P. q. qP  qq'  
            forces_mem(q,s,t2)  forces_mem(q,s,t1)},p)"
      by blast
    moreover
    assume "q  P" "qp"
    moreover
    note pP
    ultimately (*We can omit the next step but it is slower *)
    have "forces_mem(q,s,t1)  dense_below({rP. forces_mem(r,s,t2)},q)"
         "forces_mem(q,s,t2)  dense_below({rP. forces_mem(r,s,t1)},q)" 
      using aux_density_eq by simp_all
    then
    have "forces_mem(q, s, t1)  forces_mem(q, s, t2)"
      using density_mem[OF qP] by blast
  }
  with pP
  show "forces_eq(p,t1,t2)" using def_forces_eq by blast
qed

subsection‹Kunen 2013, Lemma IV.2.38›
lemma not_forces_neq:
  assumes "pP"
  shows "forces_eq(p,t1,t2)  ¬ (qP. qp  forces_neq(q,t1,t2))"
  using assms density_eq unfolding forces_neq_def by blast


lemma not_forces_nmem:
  assumes "pP"
  shows "forces_mem(p,t1,t2)  ¬ (qP. qp  forces_nmem(q,t1,t2))"
  using assms density_mem unfolding forces_nmem_def by blast


(* Use the newer versions in Forces_Definition! *)
(* (and adequate the rest of the code to them)  *)

lemma sats_forces_Nand':
  assumes
    "pP" "φformula" "ψformula" "env  list(M)" 
  shows
    "M, [p,P,leq,one] @ env  forces(Nand(φ,ψ))  
     ¬(qM. qP  is_leq(##M,leq,q,p)  
           M, [q,P,leq,one] @ env  forces(φ)  
           M, [q,P,leq,one] @ env  forces(ψ))"
  using assms sats_forces_Nand[OF assms(2-4) transitivity[OF pP]]
  P_in_M leq_in_M one_in_M unfolding forces_def
  by simp

lemma sats_forces_Neg':
  assumes
    "pP" "env  list(M)" "φformula"
  shows
    "M, [p,P,leq,one] @ env  forces(Neg(φ))    
     ¬(qM. qP  is_leq(##M,leq,q,p)  
          M, [q,P,leq,one]@env  forces(φ))"
  using assms sats_forces_Neg transitivity 
  P_in_M leq_in_M one_in_M  unfolding forces_def
  by (simp, blast)

lemma sats_forces_Forall':
  assumes
    "pP" "env  list(M)" "φformula"
  shows
    "M,[p,P,leq,one] @ env  forces(Forall(φ))  
     (xM.   M, [p,P,leq,one,x] @ env  forces(φ))"
  using assms sats_forces_Forall transitivity 
  P_in_M leq_in_M one_in_M sats_ren_forces_forall unfolding forces_def
  by simp

subsection‹The relation of forcing and atomic formulas›
lemma Forces_Equal:
  assumes
    "pP" "t1M" "t2M" "envlist(M)" "nth(n,env) = t1" "nth(m,env) = t2" "nnat" "mnat" 
  shows
    "(p  Equal(n,m) env)  forces_eq(p,t1,t2)"
   using assms sats_forces_Equal forces_eq_abs transitivity P_in_M 
  by simp

lemma Forces_Member:
  assumes
    "pP" "t1M" "t2M" "envlist(M)" "nth(n,env) = t1" "nth(m,env) = t2" "nnat" "mnat" 
  shows
    "(p  Member(n,m) env)  forces_mem(p,t1,t2)"
   using assms sats_forces_Member forces_mem_abs transitivity P_in_M
  by simp

lemma Forces_Neg:
  assumes
    "pP" "env  list(M)" "φformula" 
  shows
    "(p  Neg(φ) env)  ¬(qM. qP  qp  (q  φ env))"
    using assms sats_forces_Neg' transitivity 
  P_in_M pair_in_M_iff leq_in_M leq_abs by simp
 
subsection‹The relation of forcing and connectives›

lemma Forces_Nand:
  assumes
    "pP" "env  list(M)" "φformula" "ψformula"
  shows
    "(p  Nand(φ,ψ) env)  ¬(qM. qP  qp  (q  φ env)  (q  ψ env))"
   using assms sats_forces_Nand' transitivity 
  P_in_M pair_in_M_iff leq_in_M leq_abs by simp

lemma Forces_And_aux:
  assumes
    "pP" "env  list(M)" "φformula" "ψformula"
  shows
    "p  And(φ,ψ) env    
    (qM. qP  qp  (rM. rP  rq  (r  φ env)  (r  ψ env)))"
  unfolding And_def using assms Forces_Neg Forces_Nand by (auto simp only:)

lemma Forces_And_iff_dense_below:
  assumes
    "pP" "env  list(M)" "φformula" "ψformula"
  shows
    "(p  And(φ,ψ) env)  dense_below({rP. (r  φ env)  (r  ψ env) },p)"
  unfolding dense_below_def using Forces_And_aux assms
    by (auto dest:transitivity[OF _ P_in_M]; rename_tac q; drule_tac x=q in bspec)+

lemma Forces_Forall:
  assumes
    "pP" "env  list(M)" "φformula"
  shows
    "(p  Forall(φ) env)  (xM. (p  φ ([x] @ env)))"
   using sats_forces_Forall' assms by simp

(* "x∈val(G,π) ⟹ ∃θ. ∃p∈G.  ⟨θ,p⟩∈π ∧ val(G,θ) = x" *)
bundle some_rules =  elem_of_val_pair [dest] SepReplace_iff [simp del] SepReplace_iff[iff]

context 
  includes some_rules
begin

lemma elem_of_valI: "θ. pP. pG  θ,pπ  val(G,θ) = x  xval(G,π)"
  by (subst def_val, auto)

lemma GenExtD: "xM[G]  (τM. x = val(G,τ))"
  unfolding GenExt_def by simp

lemma left_in_M : "tauM  a,btau  aM"
  using fst_snd_closed[of "a,b"] transitivity by auto


subsection‹Kunen 2013, Lemma IV.2.29›
lemma generic_inter_dense_below: 
  assumes "DM" "M_generic(G)" "dense_below(D,p)" "pG"
  shows "D  G  0"
proof -
  let ?D="{qP. pq  qD}"
  have "dense(?D)"
  proof
    fix r
    assume "rP"
    show "d{q  P . p  q  q  D}. d  r"
    proof (cases "p  r")
      case True
      with rP
        (* Automatic tools can't handle this case for some reason... *)
      show ?thesis using leq_reflI[of r] by (intro bexI) (blast+)
    next
      case False
      then
      obtain s where "sP" "sp" "sr" by blast
      with assms rP
      show ?thesis
        using dense_belowD[OF assms(3), of s] leq_transD[of _ s r]
        by blast
    qed
  qed
  have "?DP" by auto
  (* D∈M *)
  let ?d_fm="Or(Neg(compat_in_fm(1,2,3,0)),Member(0,4))"
  have 1:"pM" 
    using M_generic(G) M_genericD transitivity[OF _ P_in_M]
          pG by simp
  moreover
  have "?d_fmformula" by simp
  moreover
  have "arity(?d_fm) = 5" unfolding compat_in_fm_def pair_fm_def upair_fm_def
    by (simp add: nat_union_abs1 Un_commute)
  moreover
  have "(M, [q,P,leq,p,D]  ?d_fm)  (¬ is_compat_in(##M,P,leq,p,q)  qD)"
    if "qM" for q
    using that sats_compat_in_fm P_in_M leq_in_M 1 DM by simp
  moreover
  have "(¬ is_compat_in(##M,P,leq,p,q)  qD)  pq  qD" if "qM" for q
    unfolding compat_def using that compat_in_abs P_in_M leq_in_M 1 by simp
  ultimately
  have "?DM" using Collect_in_M_4p[of ?d_fm _ _ _ _ _"λx y z w h. wx  xh"] 
                    P_in_M leq_in_M DM by simp
  note asm = M_generic(G) dense(?D) ?DP ?DM
  obtain x where "xG" "x?D" using M_generic_denseD[OF asm]
    by force (* by (erule bexE) does it, but the other automatic tools don't *)
  moreover from this and M_generic(G)
  have "xD"
    using M_generic_compatD[OF _ pG, of x]
      leq_reflI compatI[of _ p x] by force
  ultimately
  show ?thesis by auto
qed

subsection‹Auxiliary results for Lemma IV.2.40(a)›
lemma IV240a_mem_Collect:
  assumes
    "πM" "τM"
  shows
    "{qP. σ. r. rP  σ,r  τ  qr  forces_eq(q,π,σ)}M"
proof -
  let ?rel_pred= "λM x a1 a2 a3 a4. σ[M]. r[M]. σr[M]. 
                ra1  pair(M,σ,r,σr)  σra4  is_leq(M,a2,x,r)  is_forces_eq'(M,a1,a2,x,a3,σ)"
  let ="Exists(Exists(Exists(And(Member(1,4),And(pair_fm(2,1,0),
          And(Member(0,7),And(leq_fm(5,3,1),forces_eq_fm(4,5,3,6,2))))))))" 
  have "σM  rM" if "σ, r  τ"  for σ r
    using that τM pair_in_M_iff transitivity[of "σ,r" τ] by simp
  then
  have "?rel_pred(##M,q,P,leq,π,τ)  (σ. r. rP  σ,r  τ  qr  forces_eq(q,π,σ))"
    if "qM" for q
    unfolding forces_eq_def using assms that P_in_M leq_in_M leq_abs forces_eq'_abs pair_in_M_iff 
    by auto
  moreover
  have "(M, [q,P,leq,π,τ]  )  ?rel_pred(##M,q,P,leq,π,τ)" if "qM" for q
    using assms that sats_forces_eq'_fm sats_leq_fm P_in_M leq_in_M by simp
  moreover
  have "formula" by simp
  moreover
  have "arity()=5" 
    unfolding leq_fm_def pair_fm_def upair_fm_def
    using arity_forces_eq_fm by (simp add:nat_simp_union Un_commute)
  ultimately
  show ?thesis 
    unfolding forces_eq_def using P_in_M leq_in_M assms 
        Collect_in_M_4p[of  _ _ _ _ _ 
            "λq a1 a2 a3 a4. σ. r. ra1  σ,r  τ  qr  forces_eq'(a1,a2,q,a3,σ)"] by simp
qed

(* Lemma IV.2.40(a), membership *)
lemma IV240a_mem:
  assumes
    "M_generic(G)" "pG" "πM" "τM" "forces_mem(p,π,τ)"
    "q σ. qP  qG  σdomain(τ)  forces_eq(q,π,σ)  
      val(G,π) = val(G,σ)" (* inductive hypothesis *)
  shows
    "val(G,π)val(G,τ)"
proof (intro elem_of_valI)
  let ?D="{qP. σ. r. rP  σ,r  τ  qr  forces_eq(q,π,σ)}"
  from M_generic(G) pG
  have "pP" by blast
  moreover
  note πM τM
  ultimately
  have "?D  M" using IV240a_mem_Collect by simp
  moreover from assms pP
  have "dense_below(?D,p)"
    using forces_mem_iff_dense_below by simp
  moreover
  note M_generic(G) pG
  ultimately
  obtain q where "qG" "q?D" using generic_inter_dense_below by blast
  then
  obtain σ r where "rP" "σ,r  τ" "qr" "forces_eq(q,π,σ)" by blast
  moreover from this and qG assms
  have "r  G" "val(G,π) = val(G,σ)" by blast+
  ultimately
  show " σ. pP. p  G  σ, p  τ  val(G, σ) = val(G, π)" by auto
qed

(* Example IV.2.36 (next two lemmas) *)
lemma refl_forces_eq:"pP  forces_eq(p,x,x)"
  using def_forces_eq by simp

lemma forces_memI: "σ,rτ  pP  rP  pr  forces_mem(p,σ,τ)"
  using refl_forces_eq[of _ σ] leq_transD leq_reflI 
  by (blast intro:forces_mem_iff_dense_below[THEN iffD2])

(* Lemma IV.2.40(a), equality, first inclusion *)
lemma IV240a_eq_1st_incl:
  assumes
    "M_generic(G)" "pG" "forces_eq(p,τ,θ)"
    and
    IH:"q σ. qP  qG  σdomain(τ)  domain(θ)  
        (forces_mem(q,σ,τ)  val(G,σ)  val(G,τ)) 
        (forces_mem(q,σ,θ)  val(G,σ)  val(G,θ))"
(* Strong enough for this case: *)
(*  IH:"⋀q σ. q∈P ⟹ σ∈domain(τ) ⟹ forces_mem(q,σ,θ) ⟹ 
      val(G,σ) ∈ val(G,θ)" *)
  shows
    "val(G,τ)  val(G,θ)"
proof
  fix x
  assume "xval(G,τ)"
  then
  obtain σ r where "σ,rτ" "rG" "val(G,σ)=x" by blast
  moreover from this and pG M_generic(G)
  obtain q where "qG" "qp" "qr" by force
  moreover from this and pG M_generic(G)
  have "qP" "pP" by blast+
  moreover from calculation and M_generic(G)
  have "forces_mem(q,σ,τ)"
    using forces_memI by blast
  moreover
  note forces_eq(p,τ,θ)
  ultimately
  have "forces_mem(q,σ,θ)"
    using def_forces_eq by blast
  with qP qG IH[of q σ] σ,rτ val(G,σ) = x
  show "xval(G,θ)" by (blast)
qed

(* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *)
lemma IV240a_eq_2nd_incl:
  assumes
    "M_generic(G)" "pG" "forces_eq(p,τ,θ)"
    and
    IH:"q σ. qP  qG  σdomain(τ)  domain(θ)  
        (forces_mem(q,σ,τ)  val(G,σ)  val(G,τ)) 
        (forces_mem(q,σ,θ)  val(G,σ)  val(G,θ))"
  shows
    "val(G,θ)  val(G,τ)"
proof
  fix x
  assume "xval(G,θ)"
  then
  obtain σ r where "σ,rθ" "rG" "val(G,σ)=x" by blast
  moreover from this and pG M_generic(G)
  obtain q where "qG" "qp" "qr" by force
  moreover from this and pG M_generic(G)
  have "qP" "pP" by blast+
  moreover from calculation and M_generic(G)
  have "forces_mem(q,σ,θ)"
    using forces_memI by blast
  moreover
  note forces_eq(p,τ,θ)
  ultimately
  have "forces_mem(q,σ,τ)"
    using def_forces_eq by blast
  with qP qG IH[of q σ] σ,rθ val(G,σ) = x
  show "xval(G,τ)" by (blast)
qed

(* Lemma IV.2.40(a), equality, second inclusion--- COPY-PASTE *)
lemma IV240a_eq:
  assumes
    "M_generic(G)" "pG" "forces_eq(p,τ,θ)"
    and
    IH:"q σ. qP  qG  σdomain(τ)  domain(θ)  
        (forces_mem(q,σ,τ)  val(G,σ)  val(G,τ)) 
        (forces_mem(q,σ,θ)  val(G,σ)  val(G,θ))"
  shows
    "val(G,τ) = val(G,θ)"
  using IV240a_eq_1st_incl[OF assms] IV240a_eq_2nd_incl[OF assms] IH by blast 

subsection‹Induction on names›

lemma core_induction:
  assumes
    "τ θ p. p  P  q σ. qP ; σdomain(θ)  Q(0,τ,σ,q)  Q(1,τ,θ,p)"
    "τ θ p. p  P  q σ. qP ; σdomain(τ)  domain(θ)  Q(1,σ,τ,q)  Q(1,σ,θ,q)  Q(0,τ,θ,p)"
    "ft  2" "p  P"
  shows
    "Q(ft,τ,θ,p)"
proof -
  {
    fix ft p τ θ
    have "Transset(eclose({τ,θ}))" (is "Transset(?e)") 
      using Transset_eclose by simp
    have "τ  ?e" "θ  ?e" 
      using arg_into_eclose by simp_all
    moreover
    assume "ft  2" "p  P"
    ultimately
    have "ft,τ,θ,p 2×?e×?e×P" (is "?a2×?e×?e×P") by simp
    then 
    have "Q(ftype(?a), name1(?a), name2(?a), cond_of(?a))"
      using core_induction_aux[of ?e P Q ?a,OF Transset(?e) assms(1,2) ?a_] 
      by (clarify) (blast)
    then have "Q(ft,τ,θ,p)" by (simp add:components_simp)
  }
  then show ?thesis using assms by simp
qed

lemma forces_induction_with_conds:
  assumes
    "τ θ p. p  P  q σ. qP ; σdomain(θ)  Q(q,τ,σ)  R(p,τ,θ)"
    "τ θ p. p  P  q σ. qP ; σdomain(τ)  domain(θ)  R(q,σ,τ)  R(q,σ,θ)  Q(p,τ,θ)"
    "p  P"
  shows
    "Q(p,τ,θ)  R(p,τ,θ)"
proof -
  let ?Q="λft τ θ p. (ft = 0  Q(p,τ,θ))  (ft = 1  R(p,τ,θ))"
  from assms(1)
  have "τ θ p. p  P  q σ. qP ; σdomain(θ)  ?Q(0,τ,σ,q)  ?Q(1,τ,θ,p)"
    by simp
  moreover from assms(2)
  have "τ θ p. p  P  q σ. qP ; σdomain(τ)  domain(θ)  ?Q(1,σ,τ,q)  ?Q(1,σ,θ,q)  ?Q(0,τ,θ,p)"
    by simp
  moreover
  note pP
  ultimately
  have "?Q(ft,τ,θ,p)" if "ft2" for ft
    by (rule core_induction[OF _ _ that, of ?Q])
  then
  show ?thesis by auto
qed

lemma forces_induction:
  assumes
    "τ θ. σ. σdomain(θ)  Q(τ,σ)  R(τ,θ)"
    "τ θ. σ. σdomain(τ)  domain(θ)  R(σ,τ)  R(σ,θ)  Q(τ,θ)"
  shows
    "Q(τ,θ)  R(τ,θ)"
proof (intro forces_induction_with_conds[OF _ _ one_in_P ])
  fix τ θ p 
  assume "q  P  σ  domain(θ)  Q(τ, σ)" for q σ
  with assms(1)
  show "R(τ,θ)"
    using one_in_P by simp
next
  fix τ θ p 
    assume "q  P  σ  domain(τ)  domain(θ)  R(σ,τ)  R(σ,θ)" for q σ
    with assms(2)
    show "Q(τ,θ)"
      using one_in_P by simp
qed

subsection‹Lemma IV.2.40(a), in full›
lemma IV240a:
  assumes
    "M_generic(G)"
  shows 
    "(τM  θM  (pG. forces_eq(p,τ,θ)  val(G,τ) = val(G,θ)))  
     (τM  θM  (pG. forces_mem(p,τ,θ)  val(G,τ)  val(G,θ)))"
    (is "?Q(τ,θ)  ?R(τ,θ)")
proof (intro forces_induction[of ?Q ?R] impI)
  fix τ θ 
  assume "τM" "θM"  "σdomain(θ)  ?Q(τ,σ)" for σ
  moreover from this
  have "σdomain(θ)  forces_eq(q, τ, σ)  val(G, τ) = val(G, σ)" 
    if "qP" "qG" for q σ
    using that domain_closed[of θ] transitivity by auto
  moreover 
  note assms
  ultimately
  show "pG. forces_mem(p,τ,θ)  val(G,τ)  val(G,θ)"
    using IV240a_mem domain_closed transitivity by (simp)
next
  fix τ θ 
  assume "τM" "θM" "σ  domain(τ)  domain(θ)  ?R(σ,τ)  ?R(σ,θ)" for σ
  moreover from this
  have IH':"σ  domain(τ)  domain(θ)  qG 
            (forces_mem(q, σ, τ)  val(G, σ)  val(G, τ))  
            (forces_mem(q, σ, θ)  val(G, σ)  val(G, θ))" for q σ 
    by (auto intro:  transitivity[OF _ domain_closed[simplified]])
  ultimately
  show "pG. forces_eq(p,τ,θ)  val(G,τ) = val(G,θ)"
    using IV240a_eq[OF assms(1) _ _ IH'] by (simp)
qed

subsection‹Lemma IV.2.40(b)›
(* Lemma IV.2.40(b), membership *)
lemma IV240b_mem:
  assumes
    "M_generic(G)" "val(G,π)val(G,τ)" "πM" "τM"
    and
    IH:"σ. σdomain(τ)  val(G,π) = val(G,σ)  
      pG. forces_eq(p,π,σ)" (* inductive hypothesis *)
  shows
    "pG. forces_mem(p,π,τ)"
proof -
  from val(G,π)val(G,τ)
  obtain σ r where "rG" "σ,rτ" "val(G,π) = val(G,σ)" by auto
  moreover from this and IH
  obtain p' where "p'G" "forces_eq(p',π,σ)" by blast
  moreover
  note M_generic(G)
  ultimately
  obtain p where "pr" "pG" "forces_eq(p,π,σ)" 
    using M_generic_compatD strengthening_eq[of p'] by blast
  moreover 
  note M_generic(G)
  moreover from calculation
  have "forces_eq(q,π,σ)" if "qP" "qp" for q
    using that strengthening_eq by blast
  moreover 
  note σ,rτ rG
  ultimately
  have "rP  σ,r  τ  qr  forces_eq(q,π,σ)" if "qP" "qp" for q
    using that leq_transD[of _ p r] by blast
  then
  have "dense_below({qP. s r. rP  s,r  τ  qr  forces_eq(q,π,s)},p)"
    using leq_reflI by blast
  moreover
  note M_generic(G) pG
  moreover from calculation
  have "forces_mem(p,π,τ)" 
    using forces_mem_iff_dense_below by blast
  ultimately
  show ?thesis by blast
qed

end (* includes *)

lemma Collect_forces_eq_in_M:
  assumes "τ  M" "θ  M"
  shows "{pP. forces_eq(p,τ,θ)}  M"
  using assms Collect_in_M_4p[of "forces_eq_fm(1,2,0,3,4)" P leq τ θ 
                                  "λA x p l t1 t2. is_forces_eq(x,t1,t2)"
                                  "λ x p l t1 t2. forces_eq(x,t1,t2)" P] 
        arity_forces_eq_fm P_in_M leq_in_M sats_forces_eq_fm forces_eq_abs forces_eq_fm_type 
  by (simp add: nat_union_abs1 Un_commute)

lemma IV240b_eq_Collects:
  assumes "τ  M" "θ  M"
  shows "{pP. σdomain(τ)  domain(θ). forces_mem(p,σ,τ)  forces_nmem(p,σ,θ)}M" and
        "{pP. σdomain(τ)  domain(θ). forces_nmem(p,σ,τ)  forces_mem(p,σ,θ)}M"
proof -
  let ?rel_pred="λM x a1 a2 a3 a4. 
        σ[M]. u[M]. da3[M]. da4[M]. is_domain(M,a3,da3)  is_domain(M,a4,da4)  
          union(M,da3,da4,u)  σu  is_forces_mem'(M,a1,a2,x,σ,a3)  
          is_forces_nmem'(M,a1,a2,x,σ,a4)"
  let ="Exists(Exists(Exists(Exists(And(domain_fm(7,1),And(domain_fm(8,0),
          And(union_fm(1,0,2),And(Member(3,2),And(forces_mem_fm(5,6,4,3,7),
                              forces_nmem_fm(5,6,4,3,8))))))))))" 
  have 1:"σM" if "σ,yδ" "δM" for σ δ y
    using that pair_in_M_iff transitivity[of "σ,y" δ] by simp
  have abs1:"?rel_pred(##M,p,P,leq,τ,θ)  
        (σdomain(τ)  domain(θ). forces_mem'(P,leq,p,σ,τ)  forces_nmem'(P,leq,p,σ,θ))" 
        if "pM" for p
    unfolding forces_mem_def forces_nmem_def
    using assms that forces_mem'_abs forces_nmem'_abs P_in_M leq_in_M 
      domain_closed Un_closed 
    by (auto simp add:1[of _ _ τ] 1[of _ _ θ])
  have abs2:"?rel_pred(##M,p,P,leq,θ,τ)  (σdomain(τ)  domain(θ). 
        forces_nmem'(P,leq,p,σ,τ)  forces_mem'(P,leq,p,σ,θ))" if "pM" for p
    unfolding forces_mem_def forces_nmem_def
    using assms that forces_mem'_abs forces_nmem'_abs P_in_M leq_in_M 
      domain_closed Un_closed 
    by (auto simp add:1[of _ _ τ] 1[of _ _ θ])
  have fsats1:"(M,[p,P,leq,τ,θ]  )  ?rel_pred(##M,p,P,leq,τ,θ)" if "pM" for p
    using that assms sats_forces_mem'_fm sats_forces_nmem'_fm P_in_M leq_in_M
      domain_closed Un_closed by simp
  have fsats2:"(M,[p,P,leq,θ,τ]  )  ?rel_pred(##M,p,P,leq,θ,τ)" if "pM" for p
    using that assms sats_forces_mem'_fm sats_forces_nmem'_fm P_in_M leq_in_M
      domain_closed Un_closed by simp
  have fty:"formula" by simp
  have farit:"arity()=5"
    unfolding forces_nmem_fm_def domain_fm_def pair_fm_def upair_fm_def union_fm_def
    using arity_forces_mem_fm by (simp add:nat_simp_union Un_commute)
    show 
    "{p  P . σdomain(τ)  domain(θ). forces_mem(p, σ, τ)  forces_nmem(p, σ, θ)}  M"
    and "{p  P . σdomain(τ)  domain(θ). forces_nmem(p, σ, τ)  forces_mem(p, σ, θ)}  M"
    unfolding forces_mem_def
    using abs1 fty fsats1 farit P_in_M leq_in_M assms forces_nmem
          Collect_in_M_4p[of  _ _ _ _ _ 
          "λx p l a1 a2. (σdomain(a1)  domain(a2). forces_mem'(p,l,x,σ,a1)  
                                                     forces_nmem'(p,l,x,σ,a2))"]
    using abs2 fty fsats2 farit P_in_M leq_in_M assms forces_nmem domain_closed Un_closed
          Collect_in_M_4p[of  P leq θ τ ?rel_pred 
          "λx p l a2 a1. (σdomain(a1)  domain(a2). forces_nmem'(p,l,x,σ,a1)  
                                                     forces_mem'(p,l,x,σ,a2))" P]  
    by simp_all
qed

(* Lemma IV.2.40(b), equality *)
lemma IV240b_eq:
  assumes
    "M_generic(G)" "val(G,τ) = val(G,θ)" "τM" "θM" 
    and
    IH:"σ. σdomain(τ)domain(θ)  
      (val(G,σ)val(G,τ)  (qG. forces_mem(q,σ,τ)))  
      (val(G,σ)val(G,θ)  (qG. forces_mem(q,σ,θ)))"
    (* inductive hypothesis *)
  shows
    "pG. forces_eq(p,τ,θ)"
proof -
  let ?D1="{pP. forces_eq(p,τ,θ)}"
  let ?D2="{pP. σdomain(τ)  domain(θ). forces_mem(p,σ,τ)  forces_nmem(p,σ,θ)}"
  let ?D3="{pP. σdomain(τ)  domain(θ). forces_nmem(p,σ,τ)  forces_mem(p,σ,θ)}"
  let ?D="?D1  ?D2  ?D3"
  note assms
  moreover from this
  have "domain(τ)  domain(θ)M" (is "?BM") using domain_closed Un_closed by auto
  moreover from calculation
  have "?D2M" and "?D3M" using IV240b_eq_Collects by simp_all
  ultimately
  have "?DM" using Collect_forces_eq_in_M Un_closed by auto
  moreover
  have "dense(?D)"
  proof
    fix p
    assume "pP"
    have "dP. (forces_eq(d, τ, θ) 
            (σdomain(τ)  domain(θ). forces_mem(d, σ, τ)  forces_nmem(d, σ, θ)) 
            (σdomain(τ)  domain(θ). forces_nmem(d, σ, τ)  forces_mem(d, σ, θ))) 
           d  p" 
    proof (cases "forces_eq(p, τ, θ)")
      case True
      with pP 
      show ?thesis using leq_reflI by blast
    next
      case False
      moreover note pP
      moreover from calculation
      obtain σ q where "σdomain(τ)domain(θ)" "qP" "qp"
        "(forces_mem(q, σ, τ)  ¬ forces_mem(q, σ, θ)) 
         (¬ forces_mem(q, σ, τ)  forces_mem(q, σ, θ))"
        using def_forces_eq by blast
      moreover from this
      obtain r where "rq" "rP"
        "(forces_mem(r, σ, τ)  forces_nmem(r, σ, θ)) 
         (forces_nmem(r, σ, τ)  forces_mem(r, σ, θ))"
        using not_forces_nmem strengthening_mem by blast
      ultimately
      show ?thesis using leq_transD by blast
    qed
    then
    show "d?D1  ?D2  ?D3. d  p" by blast
  qed
  moreover
  have "?D  P"
    by auto
  moreover
  note M_generic(G)
  ultimately
  obtain p where "pG" "p?D"
    unfolding M_generic_def by blast
  then 
  consider 
    (1) "forces_eq(p,τ,θ)" | 
    (2) "σdomain(τ)  domain(θ). forces_mem(p,σ,τ)  forces_nmem(p,σ,θ)" | 
    (3) "σdomain(τ)  domain(θ). forces_nmem(p,σ,τ)  forces_mem(p,σ,θ)"
    by blast
  then
  show ?thesis
  proof (cases)
    case 1
    with pG 
    show ?thesis by blast
  next
    case 2
    then 
    obtain σ where "σdomain(τ)  domain(θ)" "forces_mem(p,σ,τ)" "forces_nmem(p,σ,θ)" 
      by blast
    moreover from this and pG and assms
    have "val(G,σ)val(G,τ)"
      using IV240a[of G σ τ] transitivity[OF _ domain_closed[simplified]] by blast
    moreover note IH val(G,τ) = _
    ultimately
    obtain q where "qG" "forces_mem(q, σ, θ)" by auto
    moreover from this and pG M_generic(G)
    obtain r where "rP" "rp" "rq"
      by blast
    moreover
    note M_generic(G)
    ultimately
    have "forces_mem(r, σ, θ)"
      using strengthening_mem by blast
    with rp forces_nmem(p,σ,θ) rP
    have "False"
      unfolding forces_nmem_def by blast
    then
    show ?thesis by simp
  next (* copy-paste from case 2 mutatis mutandis*)
    case 3
    then
    obtain σ where "σdomain(τ)  domain(θ)" "forces_mem(p,σ,θ)" "forces_nmem(p,σ,τ)" 
      by blast
    moreover from this and pG and assms
    have "val(G,σ)val(G,θ)"
      using IV240a[of G σ θ] transitivity[OF _ domain_closed[simplified]] by blast
    moreover note IH val(G,τ) = _
    ultimately
    obtain q where "qG" "forces_mem(q, σ, τ)" by auto
    moreover from this and pG M_generic(G)
    obtain r where "rP" "rp" "rq"
      by blast
    moreover
    note M_generic(G)
    ultimately
    have "forces_mem(r, σ, τ)"
      using strengthening_mem by blast
    with rp forces_nmem(p,σ,τ) rP
    have "False"
      unfolding forces_nmem_def by blast
    then
    show ?thesis by simp
  qed
qed

(* Lemma IV.2.40(b), full *)
lemma IV240b:
  assumes
    "M_generic(G)"
  shows 
    "(τMθMval(G,τ) = val(G,θ)  (pG. forces_eq(p,τ,θ))) 
     (τMθMval(G,τ)  val(G,θ)  (pG. forces_mem(p,τ,θ)))" 
    (is "?Q(τ,θ)  ?R(τ,θ)")
proof (intro forces_induction)
  fix τ θ p
  assume "σdomain(θ)  ?Q(τ, σ)" for σ
  with assms
  show "?R(τ, θ)"
    using IV240b_mem domain_closed transitivity by (simp)
next
  fix τ θ p
  assume "σ  domain(τ)  domain(θ)  ?R(σ,τ)  ?R(σ,θ)" for σ
  moreover from this
  have IH':"τM  θM  σ  domain(τ)  domain(θ) 
          (val(G, σ)  val(G, τ)  (qG. forces_mem(q, σ, τ))) 
          (val(G, σ)  val(G, θ)  (qG. forces_mem(q, σ, θ)))" for σ 
    by (blast intro:left_in_M) 
  ultimately
  show "?Q(τ,θ)"
    using IV240b_eq[OF assms(1)] by (auto)
qed

lemma map_val_in_MG:
  assumes 
    "envlist(M)"
  shows 
    "map(val(G),env)list(M[G])"
  unfolding GenExt_def using assms map_type2 by simp

lemma truth_lemma_mem:
  assumes 
    "envlist(M)" "M_generic(G)"
    "nnat" "mnat" "n<length(env)" "m<length(env)"
  shows 
    "(pG. p  Member(n,m) env)    M[G], map(val(G),env)  Member(n,m)"
  using assms IV240a[OF assms(2), of "nth(n,env)" "nth(m,env)"] 
    IV240b[OF assms(2), of "nth(n,env)" "nth(m,env)"] 
    P_in_M leq_in_M one_in_M 
    Forces_Member[of _  "nth(n,env)" "nth(m,env)" env n m] map_val_in_MG
  by (auto)

lemma truth_lemma_eq:
  assumes 
    "envlist(M)" "M_generic(G)" 
    "nnat" "mnat" "n<length(env)" "m<length(env)"
  shows 
    "(pG. p  Equal(n,m) env)    M[G], map(val(G),env)  Equal(n,m)"
  using assms IV240a(1)[OF assms(2), of "nth(n,env)" "nth(m,env)"] 
    IV240b(1)[OF assms(2), of "nth(n,env)" "nth(m,env)"] 
    P_in_M leq_in_M one_in_M 
    Forces_Equal[of _  "nth(n,env)" "nth(m,env)" env n m] map_val_in_MG
  by (auto)

lemma arities_at_aux:
  assumes
    "n  nat" "m  nat" "env  list(M)" "succ(n)  succ(m)  length(env)"
  shows
    "n < length(env)" "m < length(env)"
  using assms succ_leE[OF Un_leD1, of n "succ(m)" "length(env)"] 
   succ_leE[OF Un_leD2, of "succ(n)" m "length(env)"] by auto

subsection‹The Strenghtening Lemma›

lemma strengthening_lemma:
  assumes 
    "pP" "φformula" "rP" "rp"
  shows
    "env. envlist(M)  arity(φ)length(env)  p  φ env  r  φ env"
  using assms(2)
proof (induct)
  case (Member n m)
  then
  have "n<length(env)" "m<length(env)"
    using arities_at_aux by simp_all
  moreover
  assume "envlist(M)"
  moreover
  note assms Member
  ultimately
  show ?case 
    using Forces_Member[of _ "nth(n,env)" "nth(m,env)" env n m]
      strengthening_mem[of p r "nth(n,env)" "nth(m,env)"] by simp
next
  case (Equal n m)
  then
  have "n<length(env)" "m<length(env)"
    using arities_at_aux by simp_all
  moreover
  assume "envlist(M)"
  moreover
  note assms Equal
  ultimately
  show ?case 
    using Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m]
      strengthening_eq[of p r "nth(n,env)" "nth(m,env)"] by simp
next
  case (Nand φ ψ)
  with assms
  show ?case 
    using Forces_Nand transitivity[OF _ P_in_M] pair_in_M_iff 
      transitivity[OF _ leq_in_M] leq_transD by auto
next
  case (Forall φ)
  with assms
  have "p  φ ([x] @ env)" if "xM" for x
    using that Forces_Forall by simp
  with Forall 
  have "r  φ ([x] @ env)" if "xM" for x
    using that pred_le2 by (simp)
  with assms Forall
  show ?case 
    using Forces_Forall by simp
qed

subsection‹The Density Lemma›
lemma arity_Nand_le: 
  assumes "φ  formula" "ψ  formula" "arity(Nand(φ, ψ))  length(env)" "envlist(A)"
  shows "arity(φ)  length(env)" "arity(ψ)  length(env)"
  using assms 
  by (rule_tac Un_leD1, rule_tac [5] Un_leD2, auto)

lemma dense_below_imp_forces:
  assumes 
    "pP" "φformula"
  shows
    "env. envlist(M)  arity(φ)length(env) 
     dense_below({qP. (q  φ env)},p)  (p  φ env)"
  using assms(2)
proof (induct)
  case (Member n m)
  then
  have "n<length(env)" "m<length(env)"
    using arities_at_aux by simp_all
  moreover
  assume "envlist(M)"
  moreover
  note assms Member
  ultimately
  show ?case 
    using Forces_Member[of _ "nth(n,env)" "nth(m,env)" env n m]
      density_mem[of p "nth(n,env)" "nth(m,env)"] by simp
next
  case (Equal n m)
  then
  have "n<length(env)" "m<length(env)"
    using arities_at_aux by simp_all
  moreover
  assume "envlist(M)"
  moreover
  note assms Equal
  ultimately
  show ?case 
    using Forces_Equal[of _ "nth(n,env)" "nth(m,env)" env n m]
      density_eq[of p "nth(n,env)" "nth(m,env)"] by simp
next
case (Nand φ ψ)
  {  
    fix q
    assume "qM" "qP" "q p" "q  φ env"
    moreover 
    note Nand
    moreover from calculation
    obtain d where "dP" "d  Nand(φ, ψ) env" "d q"
      using dense_belowI by auto
    moreover from calculation
    have "¬(d ψ env)" if "d  φ env"
      using that Forces_Nand leq_reflI transitivity[OF _ P_in_M, of d] by auto
    moreover 
    note arity_Nand_le[of φ ψ]
    moreover from calculation
    have "d  φ env" 
       using strengthening_lemma[of q φ d env] Un_leD1 by auto
    ultimately
    have "¬ (q  ψ env)"
      using strengthening_lemma[of q ψ d env] by auto
  }
  with pP
  show ?case
    using Forces_Nand[symmetric, OF _ Nand(5,1,3)] by blast
next
  case (Forall φ)
  have "dense_below({qP. q  φ ([a]@env)},p)" if "aM" for a
  proof
    fix r
    assume "rP" "rp"
    with dense_below(_,p)
    obtain q where "qP" "qr" "q  Forall(φ) env"
      by blast
    moreover
    note Forall aM
    moreover from calculation
    have "q  φ ([a]@env)"
      using Forces_Forall by simp
    ultimately
    show "d  {qP. q  φ ([a]@env)}. d  P  dr"
      by auto
  qed
  moreover 
  note Forall(2)[of "Cons(_,env)"] Forall(1,3-5)
  ultimately
  have "p  φ ([a]@env)" if "aM" for a
    using that pred_le2 by simp
  with assms Forall
  show ?case using Forces_Forall by simp
qed

lemma density_lemma:
  assumes
    "pP" "φformula" "envlist(M)" "arity(φ)length(env)"
  shows
    "p  φ env      dense_below({qP. (q  φ env)},p)"
proof
  assume "dense_below({qP. (q  φ env)},p)"
  with assms
  show  "(p  φ env)"
    using dense_below_imp_forces by simp
next
  assume "p  φ env"
  with assms
  show "dense_below({qP. q  φ env},p)"
    using strengthening_lemma leq_reflI by auto
qed

subsection‹The Truth Lemma›
lemma Forces_And:
  assumes
    "pP" "env  list(M)" "φformula" "ψformula" 
    "arity(φ)  length(env)" "arity(ψ)  length(env)"
  shows
    "p  And(φ,ψ) env     (p  φ env)  (p  ψ env)"
proof
  assume "p  And(φ, ψ) env"
  with assms
  have "dense_below({r  P . (r  φ env)  (r  ψ env)}, p)"
    using Forces_And_iff_dense_below by simp
  then
  have "dense_below({r  P . (r  φ env)}, p)" "dense_below({r  P . (r  ψ env)}, p)"
    by blast+
  with assms
  show "(p  φ env)  (p  ψ env)"
    using density_lemma[symmetric] by simp
next
  assume "(p  φ env)  (p  ψ env)"
  have "dense_below({r  P . (r  φ env)  (r  ψ env)}, p)"
  proof (intro dense_belowI bexI conjI, assumption)
    fix q
    assume "qP" "q p"
    with assms (p  φ env)  (p  ψ env)
    show "q{r  P . (r  φ env)  (r  ψ env)}" "q q"
      using strengthening_lemma leq_reflI by auto
  qed
  with assms
  show "p  And(φ,ψ) env"
    using Forces_And_iff_dense_below by simp
qed

lemma Forces_Nand_alt:
  assumes
    "pP" "env  list(M)" "φformula" "ψformula" 
    "arity(φ)  length(env)" "arity(ψ)  length(env)"
  shows
    "(p  Nand(φ,ψ) env)  (p  Neg(And(φ,ψ)) env)"
  using assms Forces_Nand Forces_And Forces_Neg by auto

lemma truth_lemma_Neg:
  assumes 
    "φformula" "M_generic(G)" "envlist(M)" "arity(φ)length(env)" and
    IH: "(pG. p  φ env)  M[G], map(val(G),env)  φ"
  shows
    "(pG. p  Neg(φ) env)    M[G], map(val(G),env)  Neg(φ)"
proof (intro iffI, elim bexE, rule ccontr) 
  (* Direct implication by contradiction *)
  fix p 
  assume "pG" "p  Neg(φ) env" "¬(M[G],map(val(G),env)  Neg(φ))"
  moreover 
  note assms
  moreover from calculation
  have "M[G], map(val(G),env)  φ"
    using map_val_in_MG by simp
  with IH
  obtain r where "r  φ env" "rG" by blast
  moreover from this and M_generic(G) pG
  obtain q where "qp" "qr" "qG"
    by blast
  moreover from calculation 
  have "q  φ env"
    using strengthening_lemma[where φ=φ] by blast
  ultimately
  show "False"
    using Forces_Neg[where φ=φ] transitivity[OF _ P_in_M] by blast
next
  assume "M[G], map(val(G),env)  Neg(φ)"
  with assms 
  have "¬ (M[G], map(val(G),env)  φ)"
    using map_val_in_MG by simp
  let ?D="{pP. (p  φ env)  (p  Neg(φ) env)}"
  have "separation(##M,λp. (p  φ env))" 
      using separation_ax arity_forces assms P_in_M leq_in_M one_in_M arity_forces_le
    by simp
  moreover
  have "separation(##M,λp. (p  Neg(φ) env))"
      using separation_ax arity_forces assms P_in_M leq_in_M one_in_M arity_forces_le
    by simp
  ultimately
  have "separation(##M,λp. (p  φ env)  (p  Neg(φ) env))" 
    using separation_disj by simp
  then 
  have "?D  M" 
    using separation_closed P_in_M by simp
  moreover
  have "?D  P" by auto
  moreover
  have "dense(?D)"
  proof
    fix q
    assume "qP"
    show "d{p  P . (p  φ env)  (p  Neg(φ) env)}. d q"
    proof (cases "q  Neg(φ) env")
      case True
      with qP
      show ?thesis using leq_reflI by blast
    next
      case False
      with qP and assms
      show ?thesis using Forces_Neg by auto
    qed
  qed
  moreover
  note M_generic(G)
  ultimately
  obtain p where "pG" "(p  φ env)  (p  Neg(φ) env)"
    by blast
  then
  consider (1) "p  φ env" | (2) "p  Neg(φ) env" by blast
  then
  show "pG. (p  Neg(φ) env)"
  proof (cases)
    case 1
    with ¬ (M[G],map(val(G),env)  φ) pG IH
    show ?thesis
      by blast
  next
    case 2
    with pG 
    show ?thesis by blast
  qed
qed 

lemma truth_lemma_And:
  assumes 
    "envlist(M)" "φformula" "ψformula"
    "arity(φ)length(env)" "arity(ψ)  length(env)" "M_generic(G)"
    and
    IH: "(pG. p  φ env)     M[G], map(val(G),env)  φ"
        "(pG. p  ψ env)     M[G], map(val(G),env)  ψ"
  shows
    "(pG. (p  And(φ,ψ) env))  M[G] , map(val(G),env)  And(φ,ψ)"
  using assms map_val_in_MG Forces_And[OF M_genericD assms(1-5)]
proof (intro iffI, elim bexE)
  fix p
  assume "pG" "p  And(φ,ψ) env"
  with assms
  show "M[G], map(val(G),env)  And(φ,ψ)" 
    using Forces_And[OF M_genericD, of _ _ _ φ ψ] map_val_in_MG by auto
next 
  assume "M[G], map(val(G),env)  And(φ,ψ)"
  moreover
  note assms
  moreover from calculation
  obtain q r where "q  φ env" "r  ψ env" "qG" "rG"
    using map_val_in_MG Forces_And[OF M_genericD assms(1-5)] by auto
  moreover from calculation
  obtain p where "pq" "pr" "pG"
    by blast
  moreover from calculation
  have "(p  φ env)  (p  ψ env)" (* can't solve as separate goals *)
    using strengthening_lemma by (blast)
  ultimately
  show "pG. (p  And(φ,ψ) env)"
    using Forces_And[OF M_genericD assms(1-5)] by auto
qed 

definition 
  ren_truth_lemma :: "ii" where
  "ren_truth_lemma(φ)  
    Exists(Exists(Exists(Exists(Exists(
    And(Equal(0,5),And(Equal(1,8),And(Equal(2,9),And(Equal(3,10),And(Equal(4,6),
    iterates(λp. incr_bv(p)`5 , 6, φ)))))))))))"

lemma ren_truth_lemma_type[TC] :
  "φformula  ren_truth_lemma(φ) formula" 
  unfolding ren_truth_lemma_def
  by simp

lemma arity_ren_truth : 
  assumes "φformula"
  shows "arity(ren_truth_lemma(φ))  6  succ(arity(φ))"
proof -
  consider (lt) "5 <arity(φ)" | (ge) "¬ 5 < arity(φ)"
    by auto
  then
  show ?thesis
  proof cases
    case lt
    consider (a) "5<arity(φ)#+5"  | (b) "arity(φ)#+5  5"
      using not_lt_iff_le φ_ by force
    then 
    show ?thesis
    proof cases
      case a
      with φ_ lt
      have "5 < succ(arity(φ))" "5<arity(φ)#+2"  "5<arity(φ)#+3"  "5<arity(φ)#+4"
        using succ_ltI by auto
       with φ_ 
      have c:"arity(iterates(λp. incr_bv(p)`5,5,φ)) = 5#+arity(φ)" (is "arity(?φ') = _") 
        using arity_incr_bv_lemma lt a
        by simp
      with φ_
      have "arity(incr_bv(?φ')`5) = 6#+arity(φ)"
        using arity_incr_bv_lemma[of ?φ' 5] a by auto
      with φ_
      show ?thesis
        unfolding ren_truth_lemma_def
        using pred_Un_distrib nat_union_abs1 Un_assoc[symmetric] a c nat_union_abs2
        by simp
    next
      case b
      with φ_ lt
      have "5 < succ(arity(φ))" "5<arity(φ)#+2"  "5<arity(φ)#+3"  "5<arity(φ)#+4" "5<arity(φ)#+5"
        using succ_ltI by auto
      with φ_ 
      have "arity(iterates(λp. incr_bv(p)`5,6,φ)) = 6#+arity(φ)" (is "arity(?φ') = _") 
        using arity_incr_bv_lemma lt 
        by simp
      with φ_
      show ?thesis
        unfolding ren_truth_lemma_def
        using pred_Un_distrib nat_union_abs1 Un_assoc[symmetric]  nat_union_abs2
        by simp
    qed
  next
    case ge
    with φ_
    have "arity(φ)  5" "pred^5(arity(φ))  5"
      using not_lt_iff_le le_trans[OF le_pred]
      by auto
    with φ_
    have "arity(iterates(λp. incr_bv(p)`5,6,φ)) = arity(φ)" "arity(φ)6" "pred^5(arity(φ))  6"
      using arity_incr_bv_lemma ge le_trans[OF arity(φ)5] le_trans[OF pred^5(arity(φ))5]
      by auto
    with arity(φ)  5 φ_ pred^5(_)  5
    show ?thesis
      unfolding ren_truth_lemma_def
      using  pred_Un_distrib nat_union_abs1 Un_assoc[symmetric] nat_union_abs2 
      by simp
  qed
qed

lemma sats_ren_truth_lemma:
  "[q,b,d,a1,a2,a3] @ env  list(M)  φ  formula 
   (M, [q,b,d,a1,a2,a3] @ env  ren_truth_lemma(φ) ) 
   (M, [q,a1,a2,a3,b] @ env  φ)"
  unfolding ren_truth_lemma_def
  by (insert sats_incr_bv_iff [of _ _ M _ "[q,a1,a2,a3,b]"], simp)

lemma truth_lemma' :
  assumes
    "φformula" "envlist(M)" "arity(φ)  succ(length(env))" 
  shows
    "separation(##M,λd. bM. qP. qd  ¬(q  φ ([b]@env)))"
proof -
  let ?rel_pred="λM x a1 a2 a3. bM. qM. qa1  is_leq(##M,a2,q,x)  
                  ¬(M, [q,a1,a2,a3,b] @ env  forces(φ))" 
  let ="Exists(Forall(Implies(And(Member(0,3),leq_fm(4,0,2)),
          Neg(ren_truth_lemma(forces(φ))))))"
  have "qM" if "qP" for q using that transitivity[OF _ P_in_M] by simp
  then
  have 1:"qM. qP  R(q)  Q(q)  (qP. R(q)  Q(q))" for R Q 
    by auto
  then
  have "b  M; qM. q  P  q  d  ¬(q  φ ([b]@env)) 
         cM. qP. q  d  ¬(q  φ ([c]@env))" for b d
    by (rule bexI,simp_all)
  then
  have "?rel_pred(M,d,P,leq,one)  (bM. qP. qd  ¬(q  φ ([b]@env)))" if "dM" for d
    using that leq_abs leq_in_M P_in_M one_in_M assms
    by auto
  moreover
  have "formula" using assms by simp
  moreover
  have "(M, [d,P,leq,one]@env  )  ?rel_pred(M,d,P,leq,one)" if "dM" for d
    using assms that P_in_M leq_in_M one_in_M sats_leq_fm sats_ren_truth_lemma
    by simp
  moreover
  have "arity()  4#+length(env)" 
  proof -
    have eq:"arity(leq_fm(4, 0, 2)) = 5"
      using arity_leq_fm succ_Un_distrib nat_simp_union
      by simp
    with φ_
    have "arity() = 3  (pred^2(arity(ren_truth_lemma(forces(φ)))))"
      using nat_union_abs1 pred_Un_distrib by simp
    moreover
    have "...  3  (pred(pred(6  succ(arity(forces(φ))))))" (is "_  ?r")
      using  φ_ Un_le_compat[OF le_refl[of 3]] 
                  le_imp_subset arity_ren_truth[of "forces(φ)"]
                  pred_mono
      by auto
    finally
    have "arity()  ?r" by simp
    have i:"?r  4  pred(arity(forces(φ)))" 
      using pred_Un_distrib pred_succ_eq φ_ Un_assoc[symmetric] nat_union_abs1 by simp
    have h:"4  pred(arity(forces(φ)))  4  (4#+length(env))"
      using  env_ add_commute φ_
            Un_le_compat[of 4 4,OF _ pred_mono[OF _ arity_forces_le[OF _ _ arity(φ)_]] ]
            env_ by auto
    with φ_ env_
    show ?thesis
      using le_trans[OF  arity()  ?r  le_trans[OF i h]] nat_simp_union by simp
  qed
  ultimately
  show ?thesis using assms P_in_M leq_in_M one_in_M 
       separation_ax[of "" "[P,leq,one]@env"] 
       separation_cong[of "##M" "λy. (M, [y,P,leq,one]@env )"]
    by simp
qed


lemma truth_lemma:
  assumes 
    "φformula" "M_generic(G)"
  shows 
     "env. envlist(M)  arity(φ)length(env)  
      (pG. p  φ env)      M[G], map(val(G),env)  φ"
  using assms(1)
proof (induct)
  case (Member x y)
  then
  show ?case
    using assms truth_lemma_mem[OF envlist(M) assms(2) xnat ynat] 
      arities_at_aux by simp
next
  case (Equal x y)
  then
  show ?case
    using assms truth_lemma_eq[OF envlist(M) assms(2) xnat ynat] 
      arities_at_aux by simp
next
  case (Nand φ ψ)
  moreover 
  note M_generic(G)
  ultimately
  show ?case 
    using truth_lemma_And truth_lemma_Neg Forces_Nand_alt 
      M_genericD map_val_in_MG arity_Nand_le[of φ ψ] by auto
next
  case (Forall φ)
  with M_generic(G)
  show ?case
  proof (intro iffI)
    assume "pG. (p  Forall(φ) env)"
    with M_generic(G)
    obtain p where "pG" "pM" "pP" "p  Forall(φ) env"
      using transitivity[OF _ P_in_M] by auto
    with envlist(M) φformula
    have "p  φ ([x]@env)" if "xM" for x
      using that Forces_Forall by simp
    with pG φformula env_ arity(Forall(φ))  length(env)
      Forall(2)[of "Cons(_,env)"] 
    show "M[G], map(val(G),env)   Forall(φ)"
      using pred_le2 map_val_in_MG
      by (auto iff:GenExtD)
  next
    assume "M[G], map(val(G),env)  Forall(φ)"
    let ?D1="{dP. (d  Forall(φ) env)}"
    let ?D2="{dP. bM. qP. qd  ¬(q  φ ([b]@env))}"
    define D where "D  ?D1  ?D2"
    have arφ:"arity(φ)succ(length(env))" 
      using assms arity(Forall(φ))  length(env) φformula envlist(M) pred_le2 
      by simp
    then
    have "arity(Forall(φ))  length(env)" 
      using pred_le φformula envlist(M) by simp
    then
    have "?D1M" using Collect_forces arφ φformula envlist(M) by simp
    moreover
    have "?D2M" using envlist(M) φformula  truth_lemma' separation_closed arφ
                        P_in_M
      by simp
    ultimately
    have "DM" unfolding D_def using Un_closed by simp
    moreover
    have "D  P" unfolding D_def by auto
    moreover
    have "dense(D)" 
    proof
      fix p
      assume "pP"
      show "dD. d p"
      proof (cases "p  Forall(φ) env")
        case True
        with pP 
        show ?thesis unfolding D_def using leq_reflI by blast
      next
        case False
        with Forall pP
        obtain b where "bM" "¬(p  φ ([b]@env))"
          using Forces_Forall by blast
        moreover from this pP Forall
        have "¬dense_below({qP. q  φ ([b]@env)},p)"
          using density_lemma pred_le2  by auto
        moreover from this
        obtain d where "dp" "qP. qd  ¬(q  φ ([b] @ env))"
          "dP" by blast
        ultimately
        show ?thesis unfolding D_def by auto
      qed
    qed
    moreover
    note M_generic(G)
    ultimately
    obtain d where "d  D"  "d  G" by blast
    then
    consider (1) "d?D1" | (2) "d?D2" unfolding D_def by blast
    then
    show "pG. (p  Forall(φ) env)"
    proof (cases)
      case 1
      with dG
      show ?thesis by blast
    next
      case 2
      then
      obtain b where "bM" "qP. qd ¬(q  φ ([b] @ env))"
        by blast
      moreover from this(1) and  M[G], _   Forall(φ) and 
        Forall(2)[of "Cons(b,env)"] Forall(1,3-4) M_generic(G)
      obtain p where "pG" "pP" "p  φ ([b] @ env)" 
        using pred_le2 using map_val_in_MG by (auto iff:GenExtD)
      moreover
      note dG M_generic(G)
      ultimately
      obtain q where "qG" "qP" "qd" "qp" by blast
      moreover from this and  p  φ ([b] @ env) 
        Forall  bM pP
      have "q  φ ([b] @ env)"
        using pred_le2 strengthening_lemma by simp
      moreover 
      note qP. qd ¬(q  φ ([b] @ env))
      ultimately
      show ?thesis by simp
    qed
  qed
qed
subsection‹The ``Definition of forcing''›
lemma definition_of_forcing:
  assumes
    "pP" "φformula" "envlist(M)" "arity(φ)length(env)"
  shows
    "(p  φ env) 
     (G. M_generic(G)  pG    M[G], map(val(G),env)  φ)"
proof (intro iffI allI impI, elim conjE)
  fix G
  assume "(p  φ env)" "M_generic(G)" "p  G"
  with assms 
  show "M[G], map(val(G),env)  φ"
    using truth_lemma by blast
next
  assume 1: "G.(M_generic(G) pG) M[G] , map(val(G),env)  φ"
  {
    fix r 
    assume 2: "rP" "rp"
    then 
    obtain G where "rG" "M_generic(G)"
      using generic_filter_existence by auto
    moreover from calculation 2 pP 
    have "pG" 
      unfolding M_generic_def using filter_leqD by simp
    moreover note 1
    ultimately
    have "M[G], map(val(G),env)  φ"
      by simp
    with assms M_generic(G) 
    obtain s where "sG" "(s  φ env)"
      using truth_lemma by blast
    moreover from this and  M_generic(G) rG 
    obtain q where "qG" "qs" "qr"
      by blast
    moreover from calculation sG M_generic(G) 
    have "sP" "qP" 
      unfolding M_generic_def filter_def by auto
    moreover 
    note assms
    ultimately 
    have "qP. qr  (q  φ env)"
      using strengthening_lemma by blast
  }
  then
  have "dense_below({qP. (q  φ env)},p)"
    unfolding dense_below_def by blast
  with assms
  show "(p  φ env)"
    using density_lemma by blast
qed

lemmas definability = forces_type 
end (* forcing_data *)

end