src/HOL/Nominal/Examples/Class.thy
author urbanc
Wed, 11 Jan 2006 18:46:31 +0100
changeset 18661 dde117622dac
parent 18425 bcf13dbaa339
child 18881 c5f7cba2a675
permissions -rw-r--r--
updated to new induction principle


theory Class 
imports "../nominal" 
begin

atom_decl name coname

section {* Term-Calculus from my PHD *}

nominal_datatype trm = Ax  "name" "coname"
                 | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"  ("ImpR [_].[_]._ _" [100,100,100,100] 100)
                 | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"("ImpL [_]._ [_]._ _" [100,100,100,100,100] 100)
                 | Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"        ("Cut [_]._ [_]._" [100,100,100,100] 100)

thm trm.induct
thm trm.inducts
thm trm.induct'