# HG changeset patch # User Cezary Kaliszyk # Date 1269330122 -3600 # Node ID c69d9fb1678526d67df0766bf144f6a36b890d8b # Parent aeed597d20437207847e4ed4f20b56568ba28454 Move Ex1 and Ex2 out of Test diff -r aeed597d2043 -r c69d9fb16785 Nominal/Ex1.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex1.thy Tue Mar 23 08:42:02 2010 +0100 @@ -0,0 +1,36 @@ +theory Ex1 +imports "Parser" +begin + +text {* example 1, equivalent to example 2 from Terms *} + +atom_decl name + +ML {* val _ = recursive := false *} +nominal_datatype lam = + VAR "name" +| APP "lam" "lam" +| LAM x::"name" t::"lam" bind x in t +| LET bp::"bp" t::"lam" bind "bi bp" in t +and bp = + BP "name" "lam" +binder + bi::"bp \ atom set" +where + "bi (BP x t) = {atom x}" + +thm lam_bp.fv +thm lam_bp.supp +thm lam_bp.eq_iff +thm lam_bp.bn +thm lam_bp.perm +thm lam_bp.induct +thm lam_bp.inducts +thm lam_bp.distinct +ML {* Sign.of_sort @{theory} (@{typ lam}, @{sort fs}) *} +thm lam_bp.fv[simplified lam_bp.supp] + +end + + + diff -r aeed597d2043 -r c69d9fb16785 Nominal/Ex1rec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex1rec.thy Tue Mar 23 08:42:02 2010 +0100 @@ -0,0 +1,33 @@ +theory Ex1rec +imports "Parser" "../Attic/Prove" +begin + +atom_decl name + +ML {* val _ = recursive := true *} +nominal_datatype lam' = + VAR' "name" +| APP' "lam'" "lam'" +| LAM' x::"name" t::"lam'" bind x in t +| LET' bp::"bp'" t::"lam'" bind "bi' bp" in t +and bp' = + BP' "name" "lam'" +binder + bi'::"bp' \ atom set" +where + "bi' (BP' x t) = {atom x}" + +thm lam'_bp'.fv +thm lam'_bp'.eq_iff[no_vars] +thm lam'_bp'.bn +thm lam'_bp'.perm +thm lam'_bp'.induct +thm lam'_bp'.inducts +thm lam'_bp'.distinct +ML {* Sign.of_sort @{theory} (@{typ lam'}, @{sort fs}) *} +thm lam'_bp'.fv[simplified lam'_bp'.supp] + +end + + + diff -r aeed597d2043 -r c69d9fb16785 Nominal/Ex2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex2.thy Tue Mar 23 08:42:02 2010 +0100 @@ -0,0 +1,37 @@ +theory Ex2 +imports "Parser" +begin + +text {* example 2 *} + +atom_decl name + +ML {* val _ = recursive := false *} +nominal_datatype trm' = + Var "name" +| App "trm'" "trm'" +| Lam x::"name" t::"trm'" bind x in t +| Let p::"pat'" "trm'" t::"trm'" bind "f p" in t +and pat' = + PN +| PS "name" +| PD "name" "name" +binder + f::"pat' \ atom set" +where + "f PN = {}" +| "f (PD x y) = {atom x, atom y}" +| "f (PS x) = {atom x}" + +thm trm'_pat'.fv +thm trm'_pat'.eq_iff +thm trm'_pat'.bn +thm trm'_pat'.perm +thm trm'_pat'.induct +thm trm'_pat'.distinct +thm trm'_pat'.fv[simplified trm'_pat'.supp] + +end + + + diff -r aeed597d2043 -r c69d9fb16785 Nominal/Test.thy --- a/Nominal/Test.thy Tue Mar 23 08:33:48 2010 +0100 +++ b/Nominal/Test.thy Tue Mar 23 08:42:02 2010 +0100 @@ -1,88 +1,10 @@ theory Test -imports "Parser" "../Attic/Prove" +imports "Parser" begin atom_decl name -ML {* val _ = recursive := false *} - -text {* example 1, equivalent to example 2 from Terms *} - -nominal_datatype lam = - VAR "name" -| APP "lam" "lam" -| LAM x::"name" t::"lam" bind x in t -| LET bp::"bp" t::"lam" bind "bi bp" in t -and bp = - BP "name" "lam" -binder - bi::"bp \ atom set" -where - "bi (BP x t) = {atom x}" - -thm lam_bp.fv -thm lam_bp.supp -thm lam_bp.eq_iff -thm lam_bp.bn -thm lam_bp.perm -thm lam_bp.induct -thm lam_bp.inducts -thm lam_bp.distinct -ML {* Sign.of_sort @{theory} (@{typ lam}, @{sort fs}) *} -thm lam_bp.fv[simplified lam_bp.supp] - -ML {* val _ = recursive := true *} - -nominal_datatype lam' = - VAR' "name" -| APP' "lam'" "lam'" -| LAM' x::"name" t::"lam'" bind x in t -| LET' bp::"bp'" t::"lam'" bind "bi' bp" in t -and bp' = - BP' "name" "lam'" -binder - bi'::"bp' \ atom set" -where - "bi' (BP' x t) = {atom x}" - -thm lam'_bp'.fv -thm lam'_bp'.eq_iff[no_vars] -thm lam'_bp'.bn -thm lam'_bp'.perm -thm lam'_bp'.induct -thm lam'_bp'.inducts -thm lam'_bp'.distinct -ML {* Sign.of_sort @{theory} (@{typ lam'}, @{sort fs}) *} - -thm lam'_bp'.fv[simplified lam'_bp'.supp] - - -text {* example 2 *} - ML {* val _ = recursive := false *} -nominal_datatype trm' = - Var "name" -| App "trm'" "trm'" -| Lam x::"name" t::"trm'" bind x in t -| Let p::"pat'" "trm'" t::"trm'" bind "f p" in t -and pat' = - PN -| PS "name" -| PD "name" "name" -binder - f::"pat' \ atom set" -where - "f PN = {}" -| "f (PD x y) = {atom x, atom y}" -| "f (PS x) = {atom x}" - -thm trm'_pat'.fv -thm trm'_pat'.eq_iff -thm trm'_pat'.bn -thm trm'_pat'.perm -thm trm'_pat'.induct -thm trm'_pat'.distinct -thm trm'_pat'.fv[simplified trm'_pat'.supp] nominal_datatype trm0 = Var0 "name" @@ -444,19 +366,6 @@ | "bv8 (Bar1 v x b) = {atom v}" *) -(* Type schemes with separate datatypes *) -nominal_datatype T = - TVar "name" -| TFun "T" "T" - -(* PROBLEM: -*** exception Datatype raised -*** (line 218 of "/usr/local/src/Isabelle_16-Mar-2010/src/HOL/Tools/Datatype/datatype_aux.ML") -*** At command "nominal_datatype". -nominal_datatype TyS = - TAll xs::"name list" ty::"T" bind xs in ty -*) - (* example 9 from Peter Sewell's bestiary *) (* run out of steam at the moment *) diff -r aeed597d2043 -r c69d9fb16785 Nominal/TySch.thy --- a/Nominal/TySch.thy Tue Mar 23 08:33:48 2010 +0100 +++ b/Nominal/TySch.thy Tue Mar 23 08:42:02 2010 +0100 @@ -147,4 +147,19 @@ apply auto done +(* PROBLEM: +Type schemes with separate datatypes + +nominal_datatype T = + TVar "name" +| TFun "T" "T" +nominal_datatype TyS = + TAll xs::"name list" ty::"T" bind xs in ty + +*** exception Datatype raised +*** (line 218 of "/usr/local/src/Isabelle_16-Mar-2010/src/HOL/Tools/Datatype/datatype_aux.ML") +*** At command "nominal_datatype". +*) + + end