diff -r db0ed02eba6e -r a72a04f3d6bf Nominal/Ex/PaperTest.thy --- a/Nominal/Ex/PaperTest.thy Thu Jun 02 10:09:23 2011 +0900 +++ b/Nominal/Ex/PaperTest.thy Thu Jun 02 10:11:50 2011 +0900 @@ -28,7 +28,7 @@ fun usubst :: "utrm \ name \ name \ utrm" ("_[_:::=_]" [100,100,100] 100) where - "(UVar x)[y:::=p] = (if x=y then UnPermute x p else (UVar x))" + "(UVar x)[y:::=p] = (if x=y then UnPermute p x else (UVar x))" | "(UApp t\<^isub>1 t\<^isub>2)[y:::=p] = UApp (t\<^isub>1[y:::=p]) (t\<^isub>2[y:::=p])" | "(ULam [x].t)[y:::=p] = (if x=y then (ULam [x].t) else ULam [x].(t[y:::=p]))" | "(UConst n)[y:::=p] = UConst n" @@ -89,6 +89,29 @@ | "\(ULam [x].t)\ = Lam [x].(\t\)" | "\(UApp t1 t2)\ = App (\t1\) (\t2\)" +function + utrans :: "trm \ utrm" ("\_\" [90] 100) +where + "\Var x\ = UVar x" +| "\Const x\ = UConst x" +| "\Lam [x].t\ = ULam [x].\t\" +| "\App (App (Const ''unpermute'') (Var p)) (Var x)\ = UnPermute p x" +| "\(\p x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x)) \ \App t1 t2\ = UApp (\t1\) (\t2\)" +defer +apply(simp_all) +apply(auto)[1] +apply(case_tac x) +apply(simp_all) +apply(auto) +apply(blast) +done + +termination + apply(relation "measure (\t. size t)") + apply(simp_all) + done + + lemma elim1: "ufree t \ \t\ \(Const ''unpermute'')" apply(erule ufree.induct) @@ -101,7 +124,7 @@ apply(auto dest: elim1) done -lemma +lemma ss1: "ufree t \ uss t = ss (\t\)" apply(erule ufree.induct) apply(simp_all) @@ -111,6 +134,51 @@ apply(auto) done +lemma trans1: + shows "\\t\\ = t" +apply(induct t) +apply(simp) +apply(simp) +prefer 2 +apply(simp) +apply(case_tac "(\p x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x))") +apply(erule exE)+ +apply(simp) +apply(simp) +done + +lemma trans_subst: + shows "\t[x ::= p]\ = \t\[x :::= p]" +apply(induct rule: subst.induct) +apply(simp) +defer +apply(simp) +apply(simp) +apply(simp) +apply(case_tac "(t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x))") +apply(erule exE)+ +apply(simp only:) +apply(subst utrans.simps) +apply(subst usubst.simps) +apply(case_tac "xa = x") +apply(subst (asm) subst.simps) +apply(simp only:) +apply(subst (asm) utrans.simps) +apply(simp only: usubst.simps) +apply(simp) +apply(auto)[1] +apply(case_tac "pa = x") +apply(simp) +prefer 2 +apply(simp) +apply(simp) +done + +lemma + "ss (t[x ::= p]) = ss t" +apply(subst (2) trans1[symmetric]) +apply(subst ss1[symmetric]) + fun fr :: "trm \ name set"