Second Fixed Point Theorem
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 24 Jun 2011 10:30:06 +0900
changeset 2893 589b1a0c75e6
parent 2892 a9f3600c9ae6
child 2894 8ec94871de1e
Second Fixed Point Theorem
Nominal/Ex/SFT/Consts.thy
Nominal/Ex/SFT/Lambda.thy
Nominal/Ex/SFT/Theorem.thy
Nominal/Ex/SFT/Utils.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Ex/SFT/Consts.thy	Fri Jun 24 10:30:06 2011 +0900
@@ -0,0 +1,316 @@
+header {* Constant definitions *}
+theory Consts imports Utils Lambda begin
+
+fun Umn :: "nat \<Rightarrow> nat \<Rightarrow> lam"
+where
+  [simp del]: "Umn 0 n = \<integral>(cn 0). V (cn n)"
+| [simp del]: "Umn (Suc m) n = \<integral>(cn (Suc m)). Umn m n"
+
+lemma [simp]: "2 = Suc 1"
+  by auto
+
+lemma Lam_U:
+  "x \<noteq> y \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> z \<Longrightarrow> Umn 2 0 = \<integral>x. \<integral>y. \<integral>z. V z"
+  "x \<noteq> y \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> z \<Longrightarrow> Umn 2 1 = \<integral>x. \<integral>y. \<integral>z. V y"
+  "x \<noteq> y \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> z \<Longrightarrow> Umn 2 2 = \<integral>x. \<integral>y. \<integral>z. V x"
+  apply (simp_all add: Umn.simps Abs1_eq_iff lam.fresh fresh_at_base flip_def[symmetric] Umn.simps)
+  apply (smt Zero_not_Suc cnd flip_at_base_simps flip_at_simps)+
+  done
+
+lemma a: "n \<le> m \<Longrightarrow> atom (cn n) \<notin> supp (Umn m n)"
+  apply (induct m)
+  apply (auto simp add: lam.supp supp_at_base Umn.simps)
+  by smt
+
+lemma b: "supp (Umn m n) \<subseteq> {atom (cn n)}"
+  by (induct m) (auto simp add: lam.supp supp_at_base Umn.simps)
+
+lemma supp_U[simp]: "n \<le> m \<Longrightarrow> supp (Umn m n) = {}"
+  using a b
+  by blast
+
+lemma U_eqvt:
+  "n \<le> m \<Longrightarrow> p \<bullet> (Umn m n) = Umn m n"
+  by (rule_tac [!] perm_supp_eq) (simp_all add: fresh_star_def fresh_def)
+
+definition Var where "Var \<equiv> \<integral>cx. \<integral>cy. (V cy \<cdot> (Umn 2 2) \<cdot> V cx \<cdot> V cy)"
+definition "App \<equiv> \<integral>cx. \<integral>cy. \<integral>cz. (V cz \<cdot> Umn 2 1 \<cdot> V cx \<cdot> V cy \<cdot> V cz)"
+definition "Abs \<equiv> \<integral>cx. \<integral>cy. (V cy \<cdot> Umn 2 0 \<cdot> V cx \<cdot> V cy)"
+
+lemma Var_App_Abs:
+  "x \<noteq> e \<Longrightarrow> Var = \<integral>x. \<integral>e. (V e \<cdot> Umn 2 2 \<cdot> V x \<cdot> V e)"
+  "e \<noteq> x \<Longrightarrow> e \<noteq> y \<Longrightarrow> x \<noteq> y \<Longrightarrow> App = \<integral>x. \<integral>y. \<integral>e. (V e \<cdot> Umn 2 1 \<cdot> V x \<cdot> V y \<cdot> V e)"
+  "x \<noteq> e \<Longrightarrow> Abs = \<integral>x. \<integral>e. (V e \<cdot> Umn 2 0 \<cdot> V x \<cdot> V e)"
+  unfolding Var_def App_def Abs_def
+  by (simp_all add: Abs1_eq_iff lam.fresh flip_def[symmetric] U_eqvt fresh_def lam.supp supp_at_base)
+     (smt cx_cy_cz permute_flip_at Zero_not_Suc cnd flip_at_base_simps flip_at_simps)+
+
+lemma Var_app:
+  "Var \<cdot> x \<cdot> e \<approx> e \<cdot> Umn 2 2 \<cdot> x \<cdot> e"
+  by (rule lam2_fast_app) (simp_all add: Var_App_Abs)
+
+lemma App_app:
+  "App \<cdot> x \<cdot> y \<cdot> e \<approx> e \<cdot> Umn 2 1 \<cdot> x \<cdot> y \<cdot> e"
+  by (rule lam3_fast_app[OF Var_App_Abs(2)]) (simp_all)
+
+lemma Abs_app:
+  "Abs \<cdot> x \<cdot> e \<approx> e \<cdot> Umn 2 0 \<cdot> x \<cdot> e"
+  by (rule lam2_fast_app) (simp_all add: Var_App_Abs)
+
+lemma supp_Var_App_Abs[simp]:
+  "supp Var = {}" "supp App = {}" "supp Abs = {}"
+  by (simp_all add: Var_def App_def Abs_def lam.supp supp_at_base) blast+
+
+lemma Var_App_Abs_eqvt[eqvt]:
+  "p \<bullet> Var = Var" "p \<bullet> App = App" "p \<bullet> Abs = Abs"
+  by (rule_tac [!] perm_supp_eq) (simp_all add: fresh_star_def fresh_def)
+
+nominal_primrec
+  Numeral :: "lam \<Rightarrow> lam" ("\<lbrace>_\<rbrace>" 1000)
+where
+  "\<lbrace>V x\<rbrace> = Var \<cdot> (V x)"
+| Ap: "\<lbrace>M \<cdot> N\<rbrace> = App \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace>"
+| "\<lbrace>\<integral>x. M\<rbrace> = Abs \<cdot> (\<integral>x. \<lbrace>M\<rbrace>)"
+proof auto
+  fix x :: lam and P
+  assume "\<And>xa. x = V xa \<Longrightarrow> P" "\<And>M N. x = M \<cdot> N \<Longrightarrow> P" "\<And>xa M. x = \<integral> xa. M \<Longrightarrow> P"
+  then show "P"
+    by (rule_tac y="x" and c="0 :: perm" in lam.strong_exhaust)
+       (auto simp add: Abs1_eq_iff fresh_star_def)[3]
+next
+  fix x :: var and M and xa :: var and Ma
+  assume "[[atom x]]lst. M = [[atom xa]]lst. Ma"
+    "eqvt_at Numeral_sumC M"
+  then show "[[atom x]]lst. Numeral_sumC M = [[atom xa]]lst. Numeral_sumC Ma"
+    apply -
+    apply (erule Abs_lst1_fcb)
+    apply (simp_all add: Abs_fresh_iff)
+    apply (erule fresh_eqvt_at)
+    apply (simp_all add: finite_supp Abs1_eq_iff eqvt_at_def)
+    done
+next
+  show "eqvt Numeral_graph" unfolding eqvt_def Numeral_graph_def
+    by (rule, perm_simp, rule)
+qed
+
+termination
+  by (relation "measure (\<lambda>(t). size t)")
+     (simp_all add: lam.size)
+
+lemma numeral_eqvt[eqvt]: "p \<bullet> \<lbrace>x\<rbrace> = \<lbrace>p \<bullet> x\<rbrace>"
+  by (induct x rule: lam.induct)
+     (simp_all add: Var_App_Abs_eqvt)
+
+lemma supp_numeral[simp]:
+  "supp \<lbrace>x\<rbrace> = supp x"
+  by (induct x rule: lam.induct)
+     (simp_all add: lam.supp)
+
+lemma fresh_numeral[simp]:
+  "x \<sharp> \<lbrace>y\<rbrace> = x \<sharp> y"
+  unfolding fresh_def by simp
+
+fun app_lst :: "var \<Rightarrow> lam list \<Rightarrow> lam" where
+  "app_lst n [] = V n"
+| "app_lst n (h # t) = Ap (app_lst n t) h"
+
+lemma app_lst_eqvt[eqvt]: "p \<bullet> (app_lst t ts) = app_lst (p \<bullet> t) (p \<bullet> ts)"
+  by (induct ts arbitrary: t p) (simp_all add: eqvts)
+
+lemma supp_app_lst: "supp (app_lst x l) = {atom x} \<union> supp l"
+  apply (induct l)
+  apply (simp_all add: supp_Nil lam.supp supp_at_base supp_Cons)
+  by blast
+
+lemma app_lst_eq_iff: "app_lst n M = app_lst n N \<Longrightarrow> M = N"
+  by (induct M N rule: list_induct2') simp_all
+
+lemma app_lst_rev_eq_iff: "app_lst n (rev M) = app_lst n (rev N) \<Longrightarrow> M = N"
+  by (drule app_lst_eq_iff) simp
+
+nominal_primrec
+  Ltgt :: "lam list \<Rightarrow> lam" ("\<guillemotleft>_\<guillemotright>" 1000)
+where
+  [simp del]: "atom x \<sharp> l \<Longrightarrow> \<guillemotleft>l\<guillemotright> = \<integral>x. (app_lst x (rev l))"
+  unfolding eqvt_def Ltgt_graph_def
+  apply (rule, perm_simp, rule, rule)
+  apply (rule_tac x="x" and ?'a="var" in obtain_fresh)
+  apply (simp_all add: Abs1_eq_iff lam.fresh swap_fresh_fresh fresh_at_base)
+  apply (simp add: eqvts swap_fresh_fresh)
+  apply (case_tac "x = xa")
+  apply simp_all
+  apply (subgoal_tac "eqvt app_lst")
+  apply (erule fresh_fun_eqvt_app2)
+  apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev)
+  done
+
+termination
+  by (relation "measure (\<lambda>t. size t)")
+     (simp_all add: lam.size)
+
+lemma ltgt_eqvt[eqvt]:
+  "p \<bullet> \<guillemotleft>t\<guillemotright> = \<guillemotleft>p \<bullet> t\<guillemotright>"
+proof -
+  obtain x :: var where "atom x \<sharp> (t, p \<bullet> t)" using obtain_fresh by auto
+  then have *: "atom x \<sharp> t" "atom x \<sharp> (p \<bullet> t)" using fresh_Pair by simp_all
+  then show ?thesis using *[unfolded fresh_def]
+    apply (simp add: Abs1_eq_iff lam.fresh app_lst_eqvt Ltgt.simps)
+    apply (case_tac "p \<bullet> x = x")
+    apply (simp_all add: eqvts)
+    apply rule
+    apply (subst swap_fresh_fresh)
+    apply (simp_all add: fresh_at_base_permute_iff fresh_def[symmetric] fresh_at_base)
+    apply (subgoal_tac "eqvt app_lst")
+    apply (erule fresh_fun_eqvt_app2)
+    apply (simp_all add: fresh_at_base lam.fresh eqvt_def eqvts_raw fresh_rev)
+    done
+qed
+
+lemma ltgt_eq_iff[simp]:
+  "\<guillemotleft>M\<guillemotright> = \<guillemotleft>N\<guillemotright> \<longleftrightarrow> M = N"
+proof auto
+  obtain x :: var where "atom x \<sharp> (M, N)" using obtain_fresh by auto
+  then have *: "atom x \<sharp> M" "atom x \<sharp> N" using fresh_Pair by simp_all
+  then show "(\<guillemotleft>M\<guillemotright> = \<guillemotleft>N\<guillemotright>) \<Longrightarrow> (M = N)" by (simp add: Abs1_eq_iff app_lst_rev_eq_iff Ltgt.simps)
+qed
+
+lemma Ltgt1_app: "\<guillemotleft>[M]\<guillemotright> \<cdot> N \<approx> N \<cdot> M"
+proof -
+  obtain x :: var where "atom x \<sharp> (M, N)" using obtain_fresh by auto
+  then have "atom x \<sharp> M" "atom x \<sharp> N" using fresh_Pair by simp_all
+  then show ?thesis
+  apply (subst Ltgt.simps)
+  apply (simp add: fresh_Cons fresh_Nil)
+  apply (rule b3, rule bI, simp add: b1)
+  done
+qed
+
+lemma Ltgt3_app: "\<guillemotleft>[M,N,P]\<guillemotright> \<cdot> R \<approx> R \<cdot> M \<cdot> N \<cdot> P"
+proof -
+  obtain x :: var where "atom x \<sharp> (M, N, P, R)" using obtain_fresh by auto
+  then have *: "atom x \<sharp> (M,N,P)" "atom x \<sharp> R" using fresh_Pair by simp_all
+  then have s: "V x \<cdot> M \<cdot> N \<cdot> P [x ::= R] \<approx> R \<cdot> M \<cdot> N \<cdot> P" using b1 by simp
+  show ?thesis using *
+    apply (subst Ltgt.simps)
+  apply (simp add: fresh_Cons fresh_Nil fresh_Pair_elim)
+  apply auto[1]
+  apply (rule b3, rule bI, simp add: b1)
+  done
+qed
+
+lemma supp_ltgt[simp]:
+  "supp \<guillemotleft>t\<guillemotright> = supp t"
+proof -
+  obtain x :: var where *:"atom x \<sharp> t" using obtain_fresh by auto
+  show ?thesis using *
+  by (simp_all add: Ltgt.simps lam.supp supp_at_base supp_Nil supp_app_lst supp_rev fresh_def)
+qed
+
+lemma fresh_ltgt[simp]:
+  "x \<sharp> \<guillemotleft>[y]\<guillemotright> = x \<sharp> y"
+  "x \<sharp> \<guillemotleft>[t,r,s]\<guillemotright> = x \<sharp> (t,r,s)"
+  by (simp_all add: fresh_def supp_Cons supp_Nil supp_Pair)
+
+lemma Ltgt1_subst[simp]:
+  "\<guillemotleft>[M]\<guillemotright> [y ::= A] = \<guillemotleft>[M [y ::= A]]\<guillemotright>"
+proof -
+  obtain x :: var where a: "atom x \<sharp> (M, A, y, M [y ::= A])" using obtain_fresh by blast
+  have "x \<noteq> y" using a[simplified fresh_Pair fresh_at_base] by simp
+  then show ?thesis
+    apply (subst Ltgt.simps)
+    using a apply (simp add: fresh_Nil fresh_Cons fresh_Pair_elim)
+    apply (subst Ltgt.simps)
+    using a apply (simp add: fresh_Pair_elim fresh_Nil fresh_Cons)
+    apply (simp add: a)
+    done
+qed
+
+lemma U_app:
+  "\<guillemotleft>[A,B,C]\<guillemotright> \<cdot> Umn 2 2 \<approx> A" "\<guillemotleft>[A,B,C]\<guillemotright> \<cdot> Umn 2 1 \<approx> B" "\<guillemotleft>[A,B,C]\<guillemotright> \<cdot> Umn 2 0 \<approx> C"
+  by (rule b3, rule Ltgt3_app, rule lam3_fast_app, rule Lam_U, simp_all)
+     (rule b3, rule Ltgt3_app, rule lam3_fast_app, rule Lam_U[simplified], simp_all)+
+
+definition "F1 \<equiv> \<integral>cx. (App \<cdot> \<lbrace>Var\<rbrace> \<cdot> (Var \<cdot> V cx))"
+definition "F2 \<equiv> \<integral>cx. \<integral>cy. \<integral>cz. ((App \<cdot> (App \<cdot> \<lbrace>App\<rbrace> \<cdot> (V cz \<cdot> V cx))) \<cdot> (V cz \<cdot> V cy))"
+definition "F3 \<equiv> \<integral>cx. \<integral>cy. (App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> (\<integral>cz. (V cy \<cdot> (V cx \<cdot> V cz)))))"
+
+
+lemma Lam_F:
+  "F1 = \<integral>x. (App \<cdot> \<lbrace>Var\<rbrace> \<cdot> (Var \<cdot> V x))"
+  "a \<noteq> b \<Longrightarrow> a \<noteq> c \<Longrightarrow> c \<noteq> b \<Longrightarrow> F2 = \<integral>a. \<integral>b. \<integral>c. ((App \<cdot> (App \<cdot> \<lbrace>App\<rbrace> \<cdot> (V c \<cdot> V a))) \<cdot> (V c \<cdot> V b))"
+  "a \<noteq> b \<Longrightarrow> a \<noteq> x \<Longrightarrow> x \<noteq> b \<Longrightarrow> F3 = \<integral>a. \<integral>b. (App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> (\<integral>x. (V b \<cdot> (V a \<cdot> V x)))))"
+  apply (simp_all add: F1_def F2_def F3_def Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base)
+  apply (smt cx_cy_cz permute_flip_at)+
+  done
+
+lemma supp_F[simp]:
+  "supp F1 = {}" "supp F2 = {}" "supp F3 = {}"
+  by (simp_all add: F1_def F2_def F3_def lam.supp supp_at_base)
+     blast+
+
+lemma F_eqvt[eqvt]:
+  "p \<bullet> F1 = F1" "p \<bullet> F2 = F2" "p \<bullet> F3 = F3"
+  by (rule_tac [!] perm_supp_eq)
+     (simp_all add: fresh_star_def fresh_def)
+
+lemma F_app:
+  "F1 \<cdot> A \<approx> App \<cdot> \<lbrace>Var\<rbrace> \<cdot> (Var \<cdot> A)"
+  "F2 \<cdot> A \<cdot> B \<cdot> C \<approx> (App \<cdot> (App \<cdot> \<lbrace>App\<rbrace> \<cdot> (C \<cdot> A))) \<cdot> (C \<cdot> B)"
+  by (rule lam1_fast_app, rule Lam_F, simp_all)
+     (rule lam3_fast_app, rule Lam_F, simp_all)
+
+lemma F3_app:
+  assumes f: "atom x \<sharp> A" "atom x \<sharp> B" (* or A and B have empty support *)
+  shows "F3 \<cdot> A \<cdot> B \<approx> App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> (\<integral>x. (B \<cdot> (A \<cdot> V x))))"
+proof -
+  obtain y :: var where b: "atom y \<sharp> (x, A, B)" using obtain_fresh by blast
+  obtain z :: var where c: "atom z \<sharp> (x, y, A, B)" using obtain_fresh by blast
+  have *: "x \<noteq> z" "x \<noteq> y" "y \<noteq> z"
+    using b c by (simp_all add: fresh_Pair fresh_at_base) blast+
+  have **:
+    "atom y \<sharp> z" "atom x \<sharp> z" "atom y \<sharp> x"
+    "atom z \<sharp> y" "atom z \<sharp> x" "atom x \<sharp> y"
+    "atom x \<sharp> A" "atom y \<sharp> A" "atom z \<sharp> A"
+    "atom x \<sharp> B" "atom y \<sharp> B" "atom z \<sharp> B"
+    using b c f by (simp_all add: fresh_Pair fresh_at_base) blast+
+  show ?thesis
+    apply (simp add: Lam_F(3)[of y z x] * *[symmetric])
+    apply (rule b3) apply (rule b5) apply (rule bI)
+    apply (simp add: ** fresh_Pair * *[symmetric])
+    apply (rule b3) apply (rule bI)
+    apply (simp add: ** fresh_Pair * *[symmetric])
+    apply (rule b1)
+    done
+qed
+
+definition Lam_A1_pre : "A1 \<equiv> \<integral>cx. \<integral>cy. (F1 \<cdot> V cx)"
+definition Lam_A2_pre : "A2 \<equiv> \<integral>cx. \<integral>cy. \<integral>cz. (F2 \<cdot> V cx \<cdot> V cy \<cdot> \<guillemotleft>[V cz]\<guillemotright>)"
+definition Lam_A3_pre : "A3 \<equiv> \<integral>cx. \<integral>cy. (F3 \<cdot> V cx \<cdot> \<guillemotleft>[V cy]\<guillemotright>)"
+lemma Lam_A:
+  "x \<noteq> y \<Longrightarrow> A1 = \<integral>x. \<integral>y. (F1 \<cdot> V x)"
+  "a \<noteq> b \<Longrightarrow> a \<noteq> c \<Longrightarrow> c \<noteq> b \<Longrightarrow> A2 = \<integral>a. \<integral>b. \<integral>c. (F2 \<cdot> V a \<cdot> V b \<cdot> \<guillemotleft>[V c]\<guillemotright>)"
+  "a \<noteq> b \<Longrightarrow> A3 = \<integral>a. \<integral>b. (F3 \<cdot> V a \<cdot> \<guillemotleft>[V b]\<guillemotright>)"
+  apply (simp_all add: Lam_A1_pre Lam_A2_pre Lam_A3_pre Abs1_eq_iff lam.fresh supp_at_base Var_App_Abs_eqvt numeral_eqvt flip_def[symmetric] fresh_at_base F_eqvt ltgt_eqvt)
+  apply (smt cx_cy_cz permute_flip_at)+
+  done
+
+lemma supp_A[simp]:
+  "supp A1 = {}" "supp A2 = {}" "supp A3 = {}"
+  by (auto simp add: Lam_A1_pre Lam_A2_pre Lam_A3_pre lam.supp supp_at_base supp_Cons supp_Nil)
+
+lemma A_app:
+  "A1 \<cdot> A \<cdot> B \<approx> F1 \<cdot> A"
+  "A2 \<cdot> A \<cdot> B \<cdot> C \<approx> F2 \<cdot> A \<cdot> B \<cdot> \<guillemotleft>[C]\<guillemotright>"
+  "A3 \<cdot> A \<cdot> B \<approx> F3 \<cdot> A \<cdot> \<guillemotleft>[B]\<guillemotright>"
+  apply (rule lam2_fast_app, rule Lam_A, simp_all)
+  apply (rule lam3_fast_app, rule Lam_A, simp_all)
+  apply (rule lam2_fast_app, rule Lam_A, simp_all)
+  done
+
+definition "Num \<equiv> \<guillemotleft>[\<guillemotleft>[A1,A2,A3]\<guillemotright>]\<guillemotright>"
+
+lemma supp_Num[simp]:
+  "supp Num = {}"
+  by (auto simp only: Num_def supp_ltgt supp_Pair supp_A supp_Cons supp_Nil)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Ex/SFT/Lambda.thy	Fri Jun 24 10:30:06 2011 +0900
@@ -0,0 +1,85 @@
+header {* Definition of Lambda terms and convertibility *}
+
+theory Lambda imports Nominal2 begin
+
+lemma [simp]: "supp x = {} \<Longrightarrow> y \<sharp> x"
+  unfolding fresh_def by blast
+
+atom_decl var
+
+nominal_datatype lam =
+  V "var"
+| Ap "lam" "lam" (infixl "\<cdot>" 98)
+| Lm x::"var" l::"lam"  bind x in l ("\<integral> _. _" [97, 97] 99)
+
+nominal_primrec
+  subst :: "lam \<Rightarrow> var \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [90, 90, 90] 90)
+where
+  "(V x)[y ::= s] = (if x = y then s else (V x))"
+| "(t1 \<cdot> t2)[y ::= s] = (t1[y ::= s]) \<cdot> (t2[y ::= s])"
+| "atom x \<sharp> (y, s) \<Longrightarrow> (\<integral>x. t)[y ::= s] = \<integral>x.(t[y ::= s])"
+proof auto
+  fix a b :: lam and aa :: var and P
+  assume "\<And>x y s. a = V x \<and> aa = y \<and> b = s \<Longrightarrow> P"
+    "\<And>t1 t2 y s. a = t1 \<cdot> t2 \<and> aa = y \<and> b = s \<Longrightarrow> P"
+    "\<And>x y s t. \<lbrakk>atom x \<sharp> (y, s); a = \<integral> x. t \<and> aa = y \<and> b = s\<rbrakk> \<Longrightarrow> P"
+  then show "P"
+    by (rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust)
+       (blast, blast, simp add: fresh_star_def)
+next
+  fix x :: var and t and xa :: var and ya sa ta
+  assume *: "eqvt_at subst_sumC (t, ya, sa)"
+    "atom x \<sharp> (ya, sa)" "atom xa \<sharp> (ya, sa)"
+    "[[atom x]]lst. t = [[atom xa]]lst. ta"
+  then show "[[atom x]]lst. subst_sumC (t, ya, sa) = [[atom xa]]lst. subst_sumC (ta, ya, sa)"
+    apply -
+    apply (erule Abs_lst1_fcb)
+    apply(simp (no_asm) add: Abs_fresh_iff)
+    apply(drule_tac a="atom xa" in fresh_eqvt_at)
+    apply(simp add: finite_supp)
+    apply(simp_all add: fresh_Pair_elim Abs_fresh_iff Abs1_eq_iff)
+    apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> ya = ya")
+    apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> sa = sa")
+    apply(simp add: atom_eqvt eqvt_at_def)
+    apply(rule perm_supp_eq, simp add: supp_swap fresh_star_def fresh_Pair)+
+    done
+next
+  show "eqvt subst_graph" unfolding eqvt_def subst_graph_def
+    by (rule, perm_simp, rule)
+qed
+
+termination
+  by (relation "measure (\<lambda>(t,_,_). size t)")
+     (simp_all add: lam.size)
+
+lemma subst4[eqvt]:
+  shows "(p \<bullet> t[x ::= s]) = (p \<bullet> t)[(p \<bullet> x) ::= (p \<bullet> s)]"
+  by (induct t x s rule: subst.induct) (simp_all)
+
+lemma subst1[simp]:
+  shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
+  by (nominal_induct t avoiding: x s rule: lam.strong_induct)
+     (auto simp add: lam.fresh fresh_at_base)
+
+lemma subst2[simp]: "supp t = {} \<Longrightarrow> t[x ::= s] = t"
+  by (simp add: fresh_def)
+
+lemma subst3[simp]: "M [x ::= V x] = M"
+  by (rule_tac lam="M" and c="x" in lam.strong_induct)
+     (simp_all add: fresh_star_def lam.fresh fresh_Pair)
+
+inductive
+  beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infix "\<approx>" 80)
+where
+  bI: "(\<integral>x. M) \<cdot> N \<approx> M[x ::= N]"
+| b1: "M \<approx> M"
+| b2: "M \<approx> N \<Longrightarrow> N \<approx> M"
+| b3: "M \<approx> N \<Longrightarrow> N \<approx> L \<Longrightarrow> M \<approx> L"
+| b4: "M \<approx> N \<Longrightarrow> Z \<cdot> M \<approx> Z \<cdot> N"
+| b5: "M \<approx> N \<Longrightarrow> M \<cdot> Z \<approx> N \<cdot> Z"
+| b6: "M \<approx> N \<Longrightarrow> \<integral>x. M \<approx> \<integral>x. N"
+
+lemmas [trans] = b3
+equivariance beta
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Ex/SFT/Theorem.thy	Fri Jun 24 10:30:06 2011 +0900
@@ -0,0 +1,70 @@
+header {* The main lemma about Num and the Second Fixed Point Theorem *}
+theory Theorem imports Consts begin
+
+(*<*)
+lemmas [simp] = b3[OF bI] b1 b4 b5 supp_Num[unfolded Num_def supp_ltgt] Num_def lam.fresh[unfolded fresh_def] fresh_def b6
+lemmas app = Ltgt1_app
+(*>*)
+
+lemma Num:
+  shows "Num \<cdot> \<lbrace>M\<rbrace> \<approx> \<lbrace>\<lbrace>M\<rbrace>\<rbrace>"
+proof (induct M rule: lam.induct)
+  case (V n)
+  have "Num \<cdot> \<lbrace>V n\<rbrace> = Num \<cdot> (Var \<cdot> V n)" by simp
+  also have "... = \<guillemotleft>[\<guillemotleft>[A1,A2,A3]\<guillemotright>]\<guillemotright> \<cdot> (Var \<cdot> V n)" by simp
+  also have "... \<approx> Var \<cdot> V n \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using app .
+  also have "... \<approx> \<guillemotleft>[A1,A2,A3]\<guillemotright> \<cdot> Umn 2 2 \<cdot> V n \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using Var_app .
+  also have "... \<approx> A1 \<cdot> V n \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using U_app by simp
+  also have "... \<approx> F1 \<cdot> V n" using A_app(1) .
+  also have "... \<approx> App \<cdot> \<lbrace>Var\<rbrace> \<cdot> (Var \<cdot> V n)" using F_app(1) .
+  also have "... = \<lbrace>\<lbrace>V n\<rbrace>\<rbrace>" by simp
+  finally show "Num \<cdot> \<lbrace>V n\<rbrace> \<approx> \<lbrace>\<lbrace>V n\<rbrace>\<rbrace>".
+next
+  case (Ap M N)
+  assume IH: "Num \<cdot> \<lbrace>M\<rbrace> \<approx> \<lbrace>\<lbrace>M\<rbrace>\<rbrace>" "Num \<cdot> \<lbrace>N\<rbrace> \<approx> \<lbrace>\<lbrace>N\<rbrace>\<rbrace>"
+  have "Num \<cdot> \<lbrace>M \<cdot> N\<rbrace> = Num \<cdot> (App \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace>)" by simp
+  also have "... = \<guillemotleft>[\<guillemotleft>[A1,A2,A3]\<guillemotright>]\<guillemotright> \<cdot> (App \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace>)" by simp
+  also have "... \<approx> App \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace> \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using app .
+  also have "... \<approx> \<guillemotleft>[A1,A2,A3]\<guillemotright> \<cdot> Umn 2 1 \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace> \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using App_app .
+  also have "... \<approx> A2 \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace> \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using U_app by simp
+  also have "... \<approx> F2 \<cdot> \<lbrace>M\<rbrace> \<cdot> \<lbrace>N\<rbrace> \<cdot> Num" using A_app(2) by simp
+  also have "... \<approx> App \<cdot> (App \<cdot> \<lbrace>App\<rbrace> \<cdot> (Num \<cdot> \<lbrace>M\<rbrace>)) \<cdot> (Num \<cdot> \<lbrace>N\<rbrace>)" using F_app(2) .
+  also have "... \<approx> App \<cdot> (App \<cdot> \<lbrace>App\<rbrace> \<cdot> (\<lbrace>\<lbrace>M\<rbrace>\<rbrace>)) \<cdot> (Num \<cdot> \<lbrace>N\<rbrace>)" using IH by simp
+  also have "... \<approx> \<lbrace>\<lbrace>M \<cdot> N\<rbrace>\<rbrace>" using IH by simp
+  finally show "Num \<cdot> \<lbrace>M \<cdot> N\<rbrace> \<approx> \<lbrace>\<lbrace>M \<cdot> N\<rbrace>\<rbrace>".
+next
+  case (Lm x P)
+  assume IH: "Num \<cdot> \<lbrace>P\<rbrace> \<approx> \<lbrace>\<lbrace>P\<rbrace>\<rbrace>"
+  have "Num \<cdot> \<lbrace>\<integral> x. P\<rbrace> = Num \<cdot> (Abs \<cdot> \<integral> x. \<lbrace>P\<rbrace>)" by simp
+  also have "... = \<guillemotleft>[\<guillemotleft>[A1,A2,A3]\<guillemotright>]\<guillemotright> \<cdot> (Abs \<cdot> \<integral> x. \<lbrace>P\<rbrace>)" by simp
+  also have "... \<approx> Abs \<cdot> (\<integral> x. \<lbrace>P\<rbrace>) \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using app .
+  also have "... \<approx> \<guillemotleft>[A1,A2,A3]\<guillemotright> \<cdot> Umn 2 0 \<cdot> (\<integral> x. \<lbrace>P\<rbrace>) \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using Abs_app .
+  also have "... \<approx> A3 \<cdot> (\<integral> x. \<lbrace>P\<rbrace>) \<cdot> \<guillemotleft>[A1,A2,A3]\<guillemotright>" using U_app by simp
+  also have "... \<approx> F3 \<cdot> (\<integral> x. \<lbrace>P\<rbrace>) \<cdot> \<guillemotleft>[\<guillemotleft>[A1,A2,A3]\<guillemotright>]\<guillemotright>" using A_app(3) .
+  also have "... = F3 \<cdot> (\<integral> x. \<lbrace>P\<rbrace>) \<cdot> Num" by simp
+  also have "... \<approx> App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> \<integral> x. (Num \<cdot> ((\<integral> x. \<lbrace>P\<rbrace>) \<cdot> V x)))" by (rule F3_app) simp_all
+  also have "... \<approx> App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> \<integral> x. (Num \<cdot> \<lbrace>P\<rbrace>))" using beta_app by simp
+  also have "... \<approx> App \<cdot> \<lbrace>Abs\<rbrace> \<cdot> (Abs \<cdot> \<integral> x. \<lbrace>\<lbrace>P\<rbrace>\<rbrace>)" using IH by simp
+  also have "... = \<lbrace>\<lbrace>\<integral> x. P\<rbrace>\<rbrace>" by simp
+  finally show "Num \<cdot> \<lbrace>\<integral> x. P\<rbrace> \<approx> \<lbrace>\<lbrace>\<integral> x. P\<rbrace>\<rbrace>" .
+qed
+
+lemmas [simp] = Ap Num
+lemmas [simp del] = fresh_def Num_def
+
+theorem SFP:
+  fixes F :: lam
+  shows "\<exists>X. X \<approx> F \<cdot> \<lbrace>X\<rbrace>"
+proof -
+  obtain x :: var where [simp]:"atom x \<sharp> F" using obtain_fresh by blast
+  def W \<equiv> "\<integral>x. (F \<cdot> (App \<cdot> V x \<cdot> (Num \<cdot> V x)))"
+  def X \<equiv> "W \<cdot> \<lbrace>W\<rbrace>"
+  have a: "X = W \<cdot> \<lbrace>W\<rbrace>" unfolding X_def ..
+  also have "... = (\<integral>x. (F \<cdot> (App \<cdot> V x \<cdot> (Num \<cdot> V x)))) \<cdot> \<lbrace>W\<rbrace>" unfolding W_def ..
+  also have "... \<approx> F \<cdot> (App \<cdot> \<lbrace>W\<rbrace> \<cdot> (Num \<cdot> \<lbrace>W\<rbrace>))" by simp
+  also have "... \<approx> F \<cdot> (App \<cdot> \<lbrace>W\<rbrace> \<cdot> \<lbrace>\<lbrace>W\<rbrace>\<rbrace>)" by simp
+  also have "... \<approx> F \<cdot> \<lbrace>W \<cdot> \<lbrace>W\<rbrace>\<rbrace>" by simp
+  also have "... = F \<cdot> \<lbrace>X\<rbrace>" unfolding X_def ..
+  finally show ?thesis by blast
+qed
+(*>*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Ex/SFT/Utils.thy	Fri Jun 24 10:30:06 2011 +0900
@@ -0,0 +1,101 @@
+header {* Utilities for defining constants and functions *}
+theory Utils imports Lambda begin
+
+lemma beta_app:
+  "(\<integral> x. M) \<cdot> V x \<approx> M"
+  by (rule b3, rule bI)
+     (simp add: b1)
+
+lemma lam1_fast_app:
+  assumes leq: "\<And>a. (L = \<integral> a. (F (V a)))"
+      and su: "\<And>x. atom x \<sharp> A \<Longrightarrow> F (V x) [x ::= A] = F A"
+  shows "L \<cdot> A \<approx> F A"
+proof -
+  obtain x :: var where a: "atom x \<sharp> A" using obtain_fresh by blast
+  show ?thesis
+    by (simp add: leq[of x], rule b3, rule bI, simp add: su b1 a)
+qed
+
+lemma lam2_fast_app:
+  assumes leq: "\<And>a b. a \<noteq> b \<Longrightarrow> L = \<integral> a. \<integral> b. (F (V a) (V b))"
+     and su: "\<And>x y. atom x \<sharp> A \<Longrightarrow> atom y \<sharp> A \<Longrightarrow> atom x \<sharp> B \<Longrightarrow> atom y \<sharp> B \<Longrightarrow>
+       x \<noteq> y \<Longrightarrow> F (V x) (V y) [x ::= A] [y ::= B] = F A B"
+  shows "L \<cdot> A \<cdot> B \<approx> F A B"
+proof -
+  obtain x :: var where a: "atom x \<sharp> (A, B)" using obtain_fresh by blast
+  obtain y :: var where b: "atom y \<sharp> (x, A, B)" using obtain_fresh by blast
+  obtain z :: var where c: "atom z \<sharp> (x, y, A, B)" using obtain_fresh by blast
+  have *: "x \<noteq> y" "x \<noteq> z" "y \<noteq> z"
+    using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+
+  have ** : "atom y \<sharp> z" "atom x \<sharp> z" "atom y \<sharp> x"
+            "atom z \<sharp> y" "atom z \<sharp> x" "atom x \<sharp> y"
+            "atom x \<sharp> A" "atom y \<sharp> A" "atom z \<sharp> A"
+            "atom x \<sharp> B" "atom y \<sharp> B" "atom z \<sharp> B"
+    using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+
+  show ?thesis
+    apply (simp add: leq[OF *(1)])
+    apply (rule b3) apply (rule b5) apply (rule bI)
+    apply (simp add: ** fresh_Pair)
+    apply (rule b3) apply (rule bI) apply (simp add: su b1 ** *)
+    done
+  qed
+
+lemma lam3_fast_app:
+  assumes leq: "\<And>a b c. a \<noteq> b \<Longrightarrow> b \<noteq> c \<Longrightarrow> c \<noteq> a \<Longrightarrow>
+       L = \<integral> a. \<integral> b. \<integral> c. (F (V a) (V b) (V c))"
+     and su: "\<And>x y z. atom x \<sharp> A \<Longrightarrow> atom y \<sharp> A \<Longrightarrow> atom z \<sharp> A \<Longrightarrow>
+                     atom x \<sharp> B \<Longrightarrow> atom y \<sharp> B \<Longrightarrow> atom z \<sharp> B \<Longrightarrow>
+                     atom y \<sharp> B \<Longrightarrow> atom y \<sharp> B \<Longrightarrow> atom z \<sharp> B \<Longrightarrow>
+                     x \<noteq> y \<Longrightarrow> y \<noteq> z \<Longrightarrow> z \<noteq> x \<Longrightarrow>
+      ((F (V x) (V y) (V z))[x ::= A] [y ::= B] [z ::= C] = F A B C)"
+  shows "L \<cdot> A \<cdot> B \<cdot> C \<approx> F A B C"
+proof -
+  obtain x :: var where a: "atom x \<sharp> (A, B, C)" using obtain_fresh by blast
+  obtain y :: var where b: "atom y \<sharp> (x, A, B, C)" using obtain_fresh by blast
+  obtain z :: var where c: "atom z \<sharp> (x, y, A, B, C)" using obtain_fresh by blast
+  have *: "x \<noteq> y" "y \<noteq> z" "z \<noteq> x"
+    using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+
+  have ** : "atom y \<sharp> z" "atom x \<sharp> z" "atom y \<sharp> x"
+            "atom z \<sharp> y" "atom z \<sharp> x" "atom x \<sharp> y"
+            "atom x \<sharp> A" "atom y \<sharp> A" "atom z \<sharp> A"
+            "atom x \<sharp> B" "atom y \<sharp> B" "atom z \<sharp> B"
+            "atom x \<sharp> C" "atom y \<sharp> C" "atom z \<sharp> C"
+    using a b c by (simp_all add: fresh_Pair fresh_at_base) blast+
+  show ?thesis
+    apply (simp add: leq[OF *(1) *(2) *(3)])
+    apply (rule b3) apply (rule b5) apply (rule b5) apply (rule bI)
+    apply (simp add: ** fresh_Pair)
+    apply (rule b3) apply (rule b5) apply (rule bI)
+    apply (simp add: ** fresh_Pair)
+    apply (rule b3) apply (rule bI) apply (simp add: su b1 ** *)
+    done
+  qed
+
+definition cn :: "nat \<Rightarrow> var" where "cn n \<equiv> Abs_var (Atom (Sort ''Lambda.var'' []) n)"
+
+lemma cnd[simp]: "m \<noteq> n \<Longrightarrow> cn m \<noteq> cn n"
+  unfolding cn_def using Abs_var_inject by simp
+
+definition cx :: var where "cx \<equiv> cn 0"
+definition cy :: var where "cy \<equiv> cn 1"
+definition cz :: var where "cz \<equiv> cn 2"
+
+lemma cx_cy_cz[simp]:
+  "cx \<noteq> cy" "cx \<noteq> cz" "cz \<noteq> cy"
+  unfolding cx_def cy_def cz_def
+  by simp_all
+
+lemma noteq_fresh: "atom x \<sharp> y = (x \<noteq> y)"
+  by (simp add: fresh_at_base(2))
+
+lemma fresh_fun_eqvt_app2:
+  assumes a: "eqvt f"
+  and b: "a \<sharp> x" "a \<sharp> y"
+  shows "a \<sharp> f x y"
+  using fresh_fun_eqvt_app[OF a b(1)] a b
+  by (metis fresh_fun_app)
+
+end
+
+
+