theory Renaming imports Nat_Miscellanea "~~/src/ZF/Constructible/Formula" begin section‹Auxiliary results› lemma app_nm : "n∈nat ⟹ m∈nat ⟹ f∈n→m ⟹ x ∈ nat ⟹ f`x ∈ nat" apply(case_tac "x∈n",rule_tac m="m" in in_n_in_nat,(simp add:apply_type)+) apply(subst apply_0,subst domain_of_fun,assumption+,auto) done section‹Renaming of free variables› definition sum_id :: "[i,i] ⇒ i" where "sum_id(m,f) == λj ∈ succ(m) . if j=0 then 0 else succ(f`pred(j))" lemma sum_id0 : "sum_id(m,f)`0 = 0" by(unfold sum_id_def,simp) lemma sum_idS : "succ(x) ∈ succ(m) ⟹ sum_id(m,f)`succ(x) = succ(f`x)" by(unfold sum_id_def,simp) lemma sum_id_tc : "n ∈ nat ⟹ m ∈ nat ⟹ f ∈ n → m ⟹ sum_id(n,f) ∈ succ(n) → succ(m)" apply (rule Pi_iff [THEN iffD2],rule conjI) apply (unfold sum_id_def,rule function_lam) apply (rule conjI,auto) apply (erule_tac p="x" and A="succ(n)" and b="λ i. if i = 0 then 0 else succ(f`pred(i))" and P="x∈succ(n)×succ(m)" in lamE) apply(rename_tac j,case_tac "j=0",simp,simp add:zero_in) apply(subgoal_tac "f`pred(j) ∈ m",simp) apply(rule nat_succI,assumption+) apply (erule_tac A="n" in apply_type) apply (rule Ord_succ_mem_iff [THEN iffD1],simp) apply (subst succ_pred_eq,rule_tac A="succ(n)" in subsetD,rule naturals_subset_nat) apply (simp+) done section‹Renaming of formulas› consts ren :: "i=>i" primrec "ren(Member(x,y)) = (λ n ∈ nat . λ m ∈ nat. λf ∈ n → m. Member (f`x, f`y))" "ren(Equal(x,y)) = (λ n ∈ nat . λ m ∈ nat. λf ∈ n → m. Equal (f`x, f`y))" "ren(Nand(p,q)) = (λ n ∈ nat . λ m ∈ nat. λf ∈ n → m. Nand (ren(p)`n`m`f, ren(q)`n`m`f))" "ren(Forall(p)) = (λ n ∈ nat . λ m ∈ nat. λf ∈ n → m. Forall (ren(p)`succ(n)`succ(m)`sum_id(n,f)))" lemma arity_meml : "l ∈ nat ⟹ Member(x,y) ∈ formula ⟹ arity(Member(x,y)) ≤ l ⟹ x ∈ l" by(simp,rule subsetD,rule le_imp_subset,assumption,simp) lemma arity_memr : "l ∈ nat ⟹ Member(x,y) ∈ formula ⟹ arity(Member(x,y)) ≤ l ⟹ y ∈ l" by(simp,rule subsetD,rule le_imp_subset,assumption,simp) lemma arity_eql : "l ∈ nat ⟹ Equal(x,y) ∈ formula ⟹ arity(Equal(x,y)) ≤ l ⟹ x ∈ l" by(simp,rule subsetD,rule le_imp_subset,assumption,simp) lemma arity_eqr : "l ∈ nat ⟹ Equal(x,y) ∈ formula ⟹ arity(Equal(x,y)) ≤ l ⟹ y ∈ l" by(simp,rule subsetD,rule le_imp_subset,assumption,simp) lemma nand_ar1 : "p ∈ formula ⟹ q∈formula ⟹arity(p) ≤ arity(Nand(p,q))" by (simp,rule Un_upper1_le,simp+) lemma nand_ar2 : "p ∈ formula ⟹ q∈formula ⟹arity(q) ≤ arity(Nand(p,q))" by (simp,rule Un_upper2_le,simp+) lemma nand_ar1D : "p ∈ formula ⟹ q∈formula ⟹ arity(Nand(p,q)) ≤ n ⟹ arity(p) ≤ n" by(auto simp add: le_trans[OF Un_upper1_le[of "arity(p)" "arity(q)"]]) lemma nand_ar2D : "p ∈ formula ⟹ q∈formula ⟹ arity(Nand(p,q)) ≤ n ⟹ arity(q) ≤ n" by(auto simp add: le_trans[OF Un_upper2_le[of "arity(p)" "arity(q)"]]) lemma ren_tc : "p ∈ formula ⟹ (⋀ n m f . n ∈ nat ⟹ m ∈ nat ⟹ f ∈ n→m ⟹ ren(p)`n`m`f ∈ formula)" by (induct set:formula,auto simp add: app_nm sum_id_tc) lemma ren_arity : fixes "p" assumes "p ∈ formula" shows "⋀ n m f . n ∈ nat ⟹ m ∈ nat ⟹ f ∈ n→m ⟹ arity(p) ≤ n ⟹ arity(ren(p)`n`m`f)≤m" using assms proof (induct set:formula) case (Member x y) then have "f`x ∈ m" "f`y ∈ m" using Member assms by (simp add: arity_meml apply_funtype,simp add:arity_memr apply_funtype) then show ?case using Member by (simp add: Un_least_lt ltI) next case (Equal x y) then have "f`x ∈ m" "f`y ∈ m" using Equal assms by (simp add: arity_eql apply_funtype,simp add:arity_eqr apply_funtype) then show ?case using Equal by (simp add: Un_least_lt ltI) next case (Nand p q) then have "arity(p)≤arity(Nand(p,q))" "arity(q)≤arity(Nand(p,q))" by (subst nand_ar1,simp,simp,simp,subst nand_ar2,simp+) then have "arity(p)≤n" and "arity(q)≤n" using Nand by (rule_tac j="arity(Nand(p,q))" in le_trans,simp,simp)+ then have "arity(ren(p)`n`m`f) ≤ m" and "arity(ren(q)`n`m`f) ≤ m" using Nand by auto then show ?case using Nand by (simp add:Un_least_lt) next case (Forall p) from Forall have "succ(n)∈nat" "succ(m)∈nat" by auto from Forall have 2: "sum_id(n,f) ∈ succ(n)→succ(m)" by (simp add:sum_id_tc) from Forall have 3:"arity(p) ≤ succ(n)" by (rule_tac n="arity(p)" in natE,simp+) then have "arity(ren(p)`succ(n)`succ(m)`sum_id(n,f))≤succ(m)" using Forall ‹succ(n)∈nat› ‹succ(m)∈nat› 2 by force then show ?case using Forall 2 3 ren_tc arity_type pred_le by auto qed lemma forall_arityE : "p ∈ formula ⟹ m ∈ nat ⟹ arity(Forall(p)) ≤ m ⟹ arity(p) ≤ succ(m)" by(rule_tac n="arity(p)" in natE,erule arity_type,simp+) lemma env_coincidence_sum_id : assumes "m ∈ nat" "n ∈ nat" "ρ ∈ list(A)" "ρ' ∈ list(A)" "f ∈ n → m" "⋀ i . i < n ⟹ nth(i,ρ) = nth(f`i,ρ')" "a ∈ A" "j ∈ succ(n)" shows "nth(j,Cons(a,ρ)) = nth(sum_id(n,f)`j,Cons(a,ρ'))" proof - let ?g="sum_id(n,f)" have "succ(n) ∈ nat" using ‹n∈nat› by simp then have "j ∈ nat" using ‹j∈succ(n)› in_n_in_nat by blast then have "nth(j,Cons(a,ρ)) = nth(?g`j,Cons(a,ρ'))" proof (cases rule:natE[OF ‹j∈nat›]) case 1 then show ?thesis using assms sum_id0 by simp next case (2 i) with ‹j∈succ(n)› have "succ(i)∈succ(n)" by simp with ‹n∈nat› have "i ∈ n" using nat_succD assms by simp have "f`i∈m" using ‹f∈n→m› apply_type ‹i∈n› by simp then have "f`i ∈ nat" using in_n_in_nat ‹m∈nat› by simp have "nth(succ(i),Cons(a,ρ)) = nth(i,ρ)" using ‹i∈nat› by simp also have "... = nth(f`i,ρ')" using assms ‹i∈n› ltI by simp also have "... = nth(succ(f`i),Cons(a,ρ'))" using ‹f`i∈nat› by simp also have "... = nth(?g`succ(i),Cons(a,ρ'))" using sum_idS ‹succ(i)∈succ(n)› cases by simp finally have "nth(succ(i),Cons(a,ρ)) = nth(?g`succ(i),Cons(a,ρ'))" . then show ?thesis using ‹j=succ(i)› by simp qed then show ?thesis . qed lemma sats_iff_sats_ren : fixes "φ" assumes "φ ∈ formula" shows "⟦ n ∈ nat ; m ∈ nat ; ρ ∈ list(M) ; ρ' ∈ list(M) ; f ∈ n → m ; arity(φ) ≤ n ; ⋀ i . i < n ⟹ nth(i,ρ) = nth(f`i,ρ') ⟧ ⟹ sats(M,φ,ρ) ⟷ sats(M,ren(φ)`n`m`f,ρ')" using ‹φ ∈ formula› proof(induct φ arbitrary:n m ρ ρ' f) case (Member x y) have 0: "ren(Member(x,y))`n`m`f = Member(f`x,f`y)" using Member assms arity_type by force have 1: "x ∈ n" using Member arity_meml by simp have "y ∈ n" using Member arity_memr by simp then show ?case using Member 1 0 ltI by simp next case (Equal x y) have 0: "ren(Equal(x,y))`n`m`f = Equal(f`x,f`y)" using Equal assms arity_type by force have 1: "x ∈ n" using Equal arity_eql by simp have "y ∈ n" using Equal arity_eqr by simp then show ?case using Equal 1 0 ltI by simp next case (Nand p q) have 0:"ren(Nand(p,q))`n`m`f = Nand(ren(p)`n`m`f,ren(q)`n`m`f)" using Nand by simp have "arity(p) ≤ n" using Nand nand_ar1D by simp then have 1:"i ∈ arity(p) ⟹ i ∈ n" for i using subsetD[OF le_imp_subset[OF ‹arity(p) ≤ n›]] by simp then have "i ∈ arity(p) ⟹ nth(i,ρ) = nth(f`i,ρ')" for i using Nand ltI by simp then have 2:"sats(M,p,ρ) ⟷ sats(M,ren(p)`n`m`f,ρ')" using ‹arity(p)≤n› 1 Nand by simp have "arity(q) ≤ n" using Nand nand_ar2D by simp then have 3:"i ∈ arity(q) ⟹ i ∈ n" for i using subsetD[OF le_imp_subset[OF ‹arity(q) ≤ n›]] by simp then have "i ∈ arity(q) ⟹ nth(i,ρ) = nth(f`i,ρ')" for i using Nand ltI by simp then have 4:"sats(M,q,ρ) ⟷ sats(M,ren(q)`n`m`f,ρ')" using assms ‹arity(q)≤n› 3 Nand by simp then show ?case using Nand 0 2 4 by simp next case (Forall p) have 0:"ren(Forall(p))`n`m`f = Forall(ren(p)`succ(n)`succ(m)`sum_id(n,f))" using Forall by simp have 1:"sum_id(n,f) ∈ succ(n) → succ(m)" (is "?g ∈ _") using sum_id_tc Forall by simp then have 2: "arity(p) ≤ succ(n)" using Forall le_trans[of _ "succ(pred(arity(p)))"] succpred_leI by simp have "succ(n)∈nat" "succ(m)∈nat" using Forall by auto then have A:"⋀ j .j < succ(n) ⟹ nth(j, Cons(a, ρ)) = nth(?g`j, Cons(a, ρ'))" if "a∈M" for a using that env_coincidence_sum_id Forall ltD by force have 4: "sats(M,p,Cons(a,ρ)) ⟷ sats(M,ren(p)`succ(n)`succ(m)`?g,Cons(a,ρ'))" if "a∈M" for a proof - have C:"Cons(a,ρ) ∈ list(M)" "Cons(a,ρ')∈list(M)" using Forall that by auto have "sats(M,p,Cons(a,ρ)) ⟷ sats(M,ren(p)`succ(n)`succ(m)`?g,Cons(a,ρ'))" using Forall(2)[OF ‹succ(n)∈nat› ‹succ(m)∈nat› C(1) C(2) 1 2 A[OF ‹a∈M›]] by simp then show ?thesis . qed then show ?case using Forall 0 1 2 4 by simp qed end