diff -r 2cb34b1e7e19 -r 8412c7e503d4 Nominal/Ex/PaperTest.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/PaperTest.thy Fri May 13 14:50:17 2011 +0100 @@ -0,0 +1,235 @@ +theory PaperTest +imports "../Nominal2" +begin + +atom_decl name + +datatype trm = + Const "string" +| Var "name" +| App "trm" "trm" +| Lam "name" "trm" ("Lam [_]. _" [100, 100] 100) + +fun + subst :: "trm \ name \ name \ trm" ("_[_::=_]" [100,100,100] 100) +where + "(Var x)[y::=p] = (if x=y then (App (App (Const ''unpermute'') (Var p)) (Var x)) else (Var x))" +| "(App t\<^isub>1 t\<^isub>2)[y::=p] = App (t\<^isub>1[y::=p]) (t\<^isub>2[y::=p])" +| "(Lam [x].t)[y::=p] = (if x=y then (Lam [x].t) else Lam [x].(t[y::=p]))" +| "(Const n)[y::=p] = Const n" + +datatype utrm = + UConst "string" +| UnPermute "name" "name" +| UVar "name" +| UApp "utrm" "utrm" +| ULam "name" "utrm" ("ULam [_]. _" [100, 100] 100) + +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))" +| "(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" +| "(UnPermute x q)[y:::=p] = UnPermute x q" + +function + ss :: "trm \ nat" +where + "ss (Var x) = 1" +| "ss (Const s) = 1" +| "ss (Lam [x].t) = 1 + ss t" +| "ss (App (App (Const ''unpermute'') (Var p)) (Var x)) = 1" +| "\(\p x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x)) \ ss (App t1 t2) = 1 + ss t1 + ss 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 + +fun + uss :: "utrm \ nat" +where + "uss (UVar x) = 1" +| "uss (UConst s) = 1" +| "uss (ULam [x].t) = 1 + uss t" +| "uss (UnPermute x y) = 1" +| "uss (UApp t1 t2) = 1 + uss t1 + uss t2" + +lemma trm_uss: + shows "uss (t[x:::=p]) = uss t" +apply(induct rule: uss.induct) +apply(simp_all) +done + +inductive + ufree :: "utrm \ bool" +where + "ufree (UVar x)" +| "s \ ''unpermute'' \ ufree (UConst s)" +| "ufree t \ ufree (ULam [x].t)" +| "ufree (UnPermute x y)" +| "\ufree t1; ufree t2\ \ ufree (UApp t1 t2)" + +fun + trans :: "utrm \ trm" ("\_\" [100] 100) +where + "\(UVar x)\ = Var x" +| "\(UConst x)\ = Const x" +| "\(UnPermute p x)\ = (App (App (Const ''unpermute'') (Var p)) (Var x))" +| "\(ULam [x].t)\ = Lam [x].(\t\)" +| "\(UApp t1 t2)\ = App (\t1\) (\t2\)" + +lemma elim1: + "ufree t \ \t\ \(Const ''unpermute'')" +apply(erule ufree.induct) +apply(auto) +done + +lemma elim2: + "ufree t \ \(\p. \t\ = App (Const ''unpermute'') (Var p))" +apply(erule ufree.induct) +apply(auto dest: elim1) +done + +lemma + "ufree t \ uss t = ss (\t\)" +apply(erule ufree.induct) +apply(simp_all) +apply(subst ss.simps) +apply(auto) +apply(drule elim2) +apply(auto) +done + + +fun + fr :: "trm \ name set" +where + "fr (Var x) = {x}" +| "fr (Const s) = {}" +| "fr (Lam [x].t) = fr t - {x}" +| "fr (App t1 t2) = fr t1 \ fr t2" + +function + sfr :: "trm \ name set" +where + "sfr (Var x) = {}" +| "sfr (Const s) = {}" +| "sfr (Lam [x].t) = sfr t - {x}" +| "sfr (App (App (Const ''unpermute'') (Var p)) (Var x)) = {p, x}" +| "\(\p x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x)) \ sfr (App t1 t2) = sfr t1 \ sfr 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 + +function + ss :: "trm \ nat" +where + "ss (Var x) = 1" +| "ss (Const s) = 1" +| "ss (Lam [x].t) = 1 + ss t" +| "ss (App (App (Const ''unpermute'') (Var p)) (Var x)) = 1" +| "\(\p x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x)) \ ss (App t1 t2) = 1 + ss t1 + ss 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 + +inductive + improper :: "trm \ bool" +where + "improper (App (App (Const ''unpermute'') (Var p)) (Var x))" +| "improper p x t \ improper p x (Lam [y].t)" +| "\improper p x t1; improper p x t2\ \ improper p x (App t1 t2)" + +lemma trm_ss: + shows "\improper p x t \ ss (t[x::= p]) = ss t" +apply(induct rule: ss.induct) +apply(simp) +apply(simp) +apply(simp) +apply(auto)[1] +apply(case_tac "improper p x t") +apply(drule improper.intros(2)) +apply(blast) +apply(simp) +using improper.intros(1) +apply(blast) +apply(erule contrapos_pn) +thm contrapos_np +apply(simp) +apply(auto)[1] +apply(simp) +apply(erule disjE) +apply(erule conjE)+ +apply(subst ss.simps) +apply(simp) +apply(rule disjI1) +apply(rule allI) +apply(rule notI) + +apply(simp del: ss.simps) +apply(erule disjE) +apply(subst ss.simps) +apply(simp) +apply(simp only: subst.simps) +apply(subst ss.simps) +apply(simp del: ss.simps) +apply(rule conjI) +apply(rule impI) +apply(erule conjE) +apply(erule exE)+ +apply(subst ss.simps) +apply(simp) +apply(auto)[1] +apply(simp add: if_splits) +apply() + +function + simp :: "name \ trm \ trm" +where + "simp p (Const c) = (Const c)" +| "simp p (Var x) = App (App (Const ''permute'') (Var p)) (Var x)" +| "simp p (App t1 t2) = (if ((\x. t1 = App (Const ''unpermute'') (Var p) \ t2 = (Var x))) + then t2 + else App (simp p t1) (simp p t2))" +| "simp p (Lam [x].t) = Lam [x]. (simp p (t[x ::= p]))" +apply(pat_completeness) +apply(simp_all) +apply(auto) +done + +termination +apply(relation "measure (\(_, t). size t)") +apply(simp_all) + +end \ No newline at end of file