src/HOL/Nominal/Examples/Class.thy
author urbanc
Sun, 19 Feb 2006 17:18:39 +0100
changeset 19107 b16a45c53884
parent 18881 c5f7cba2a675
child 19319 7e1f85ceb1a2
permissions -rw-r--r--
added a few lemmas to do with permutation-equivalence for the recursion combinator
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     1
18661
dde117622dac updated to new induction principle
urbanc
parents: 18425
diff changeset
     2
theory Class 
18425
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     3
imports "../nominal" 
bcf13dbaa339 I think the earlier version was completely broken
urbanc
parents: 18395
diff changeset
     4
begin
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     5
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     6
section {* Term-Calculus from Urban's PhD *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
     7
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     8
atom_decl name coname
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
     9
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    10
nominal_datatype trm = 
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    11
    Ax   "name" "coname"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    12
  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ (_)._" [100,100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    13
  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR (_)._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    14
  | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 ("NotL <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    15
  | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    16
  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 (_)._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    17
  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 (_)._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    18
  | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR1 <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    19
  | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR2 <_>._ _" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    20
  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL (_)._ (_)._ _" [100,100,100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    21
  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR (_).<_>._ _" [100,100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    22
  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    23
  
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    24
  
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    25
text {* Induction principles *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    26
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    27
thm trm.induct_weak --"weak"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    28
thm trm.induct      --"strong"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    29
thm trm.induct'     --"strong with explicit context (rarely needed)"
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    30
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    31
text {* named terms *}
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    32
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    33
nominal_datatype ntrm = N "\<guillemotleft>name\<guillemotright>trm" ("N (_)._" [100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    34
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    35
text {* conamed terms *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    36
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    37
nominal_datatype ctrm = C "\<guillemotleft>coname\<guillemotright>trm" ("C <_>._" [100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    38
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    39
text {* We should now define the two forms of substitition :o( *}
18395
87217764cec2 initial commit (not to be seen by the public)
urbanc
parents:
diff changeset
    40
18881
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    41
consts
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    42
  substn :: "trm \<Rightarrow> name   \<Rightarrow> ctrm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100) 
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    43
  substc :: "trm \<Rightarrow> coname \<Rightarrow> ntrm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    44
c5f7cba2a675 added all constructors from PhD
urbanc
parents: 18661
diff changeset
    45
text {* does not work yet *}