diff -r b8a1c3cdf739 -r c5f7cba2a675 src/HOL/Nominal/Examples/Class.thy --- a/src/HOL/Nominal/Examples/Class.thy Tue Jan 31 18:19:36 2006 +0100 +++ b/src/HOL/Nominal/Examples/Class.thy Wed Feb 01 01:03:41 2006 +0100 @@ -3,16 +3,43 @@ imports "../nominal" begin +section {* Term-Calculus from Urban's PhD *} + atom_decl name coname -section {* Term-Calculus from my PHD *} +nominal_datatype trm = + Ax "name" "coname" + | Cut "\coname\trm" "\name\trm" ("Cut <_>._ (_)._" [100,100,100,100] 100) + | NotR "\name\trm" "coname" ("NotR (_)._ _" [100,100,100] 100) + | NotL "\coname\trm" "name" ("NotL <_>._ _" [100,100,100] 100) + | AndR "\coname\trm" "\coname\trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100) + | AndL1 "\name\trm" "name" ("AndL1 (_)._ _" [100,100,100] 100) + | AndL2 "\name\trm" "name" ("AndL2 (_)._ _" [100,100,100] 100) + | OrR1 "\coname\trm" "coname" ("OrR1 <_>._ _" [100,100,100] 100) + | OrR2 "\coname\trm" "coname" ("OrR2 <_>._ _" [100,100,100] 100) + | OrL "\name\trm" "\name\trm" "name" ("OrL (_)._ (_)._ _" [100,100,100,100,100] 100) + | ImpR "\name\(\coname\trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 100) + | ImpL "\coname\trm" "\name\trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100) + + +text {* Induction principles *} -nominal_datatype trm = Ax "name" "coname" - | ImpR "\name\(\coname\trm)" "coname" ("ImpR [_].[_]._ _" [100,100,100,100] 100) - | ImpL "\coname\trm" "\name\trm" "name"("ImpL [_]._ [_]._ _" [100,100,100,100,100] 100) - | Cut "\coname\trm" "\name\trm" ("Cut [_]._ [_]._" [100,100,100,100] 100) +thm trm.induct_weak --"weak" +thm trm.induct --"strong" +thm trm.induct' --"strong with explicit context (rarely needed)" + +text {* named terms *} + +nominal_datatype ntrm = N "\name\trm" ("N (_)._" [100,100] 100) + +text {* conamed terms *} -thm trm.induct -thm trm.inducts -thm trm.induct' +nominal_datatype ctrm = C "\coname\trm" ("C <_>._" [100,100] 100) + +text {* We should now define the two forms of substitition :o( *} +consts + substn :: "trm \ name \ ctrm \ trm" ("_[_::=_]" [100,100,100] 100) + substc :: "trm \ coname \ ntrm \ trm" ("_[_::=_]" [100,100,100] 100) + +text {* does not work yet *} \ No newline at end of file