# HG changeset patch # User Cezary Kaliszyk # Date 1306977110 -32400 # Node ID a72a04f3d6bf49b30a5662b9e80ed0f7df00ab3d # Parent db0ed02eba6ec3ab9f95712d0b2cdc4c91bd60c0# Parent 04f7c4ce8588a08a65e8159532c4cf05d42af20f merge diff -r db0ed02eba6e -r a72a04f3d6bf Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Thu Jun 02 10:09:23 2011 +0900 +++ b/Nominal/Ex/Lambda.thy Thu Jun 02 10:11:50 2011 +0900 @@ -10,6 +10,7 @@ | App "lam" "lam" | Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100) + inductive triv :: "lam \ nat \ bool" where @@ -451,10 +452,12 @@ apply auto done +(* lemma db_trans_test: assumes a: "y \ x" shows "trans (Lam [x]. Lam [y]. App (Var x) (Var y)) [] 0 = Some (DBLam (DBLam (DBApp (DBVar 1) (DBVar 0))))" using a by simp +*) abbreviation mbind :: "'a option => ('a => 'b option) => 'b option" ("_ \= _" [65,65] 65) @@ -589,8 +592,6 @@ "Z (App M N) k = Z M (%m. (Z N (%n.(App m n))))" unfolding eqvt_def Z_graph_def apply (rule, perm_simp, rule) -prefer 2 -ML_prf {* prop_of (hd (#prems (fst (Subgoal.focus @{context} 1 (#goal (Isar.goal ())))))) *} oops (* function tests *) 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" diff -r db0ed02eba6e -r a72a04f3d6bf Nominal/Ex/TypeSchemes.thy diff -r db0ed02eba6e -r a72a04f3d6bf Nominal/nominal_function_core.ML --- a/Nominal/nominal_function_core.ML Thu Jun 02 10:09:23 2011 +0900 +++ b/Nominal/nominal_function_core.ML Thu Jun 02 10:11:50 2011 +0900 @@ -124,9 +124,10 @@ end (** building proof obligations *) -fun mk_eqvt_proof_obligation qs fvar (_, assms, arg) = +fun mk_eqvt_proof_obligation qs fvar (vs, assms, arg) = mk_eqvt_at (fvar, arg) |> curry Logic.list_implies (map prop_of assms) + |> curry Term.list_all_free vs |> curry Term.list_abs_free qs |> strip_abs_body @@ -336,7 +337,7 @@ fun prep_eqvt (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_eqvt_case) |> fold_rev (Thm.implies_intr o cprop_of) CCas - (* fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs *) + |> fold_rev (Thm.forall_intr o cterm_of thy o Free) RIvs in map prep_eqvt RCs |> map (fold_rev (Thm.implies_intr o cprop_of) ags)