Nominal/Ex/Shallow.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 15 Nov 2010 08:17:11 +0000
changeset 2570 1c77e15c4259
child 2622 e6e6a3da81aa
permissions -rw-r--r--
added a test for the various shallow binders
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2570
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Shallow
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "../Nominal2" 
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
(* 
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
  Various shallow binders
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
*)
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
atom_decl name
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
text {* binding lists of names *}
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
nominal_datatype trm1 =
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  Var1 "name"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| App1 "trm1" "trm1"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
| Lam1 xs::"name list" t::"trm1" bind xs in t
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
text {* binding a finite set of names *}
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
nominal_datatype trm2 =
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  Var2 "name"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
| App2 "trm2" "trm2"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
| Lam2 xs::"name fset" t::"trm2" bind (set) xs in t
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
text {* restricting a finite set of names *}
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
nominal_datatype trm3 =
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
  Var3 "name"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
| App3 "trm3" "trm3"
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
| Lam3 xs::"name fset" t::"trm3" bind (res) xs in t
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
end
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
1c77e15c4259 added a test for the various shallow binders
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38