Nominal/Ex/Shallow.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 28 Dec 2010 00:20:50 +0000
changeset 2629 ffb5a181844b
parent 2622 e6e6a3da81aa
child 2634 3ce1089cdbf3
permissions -rw-r--r--
proper application of induction_schema and strong_exhaust rules; needs local fix in induction_schema.ML

theory Shallow
imports "../Nominal2" 
begin

(* 
  Various shallow binders
*)

atom_decl name

text {* binding lists of names *}

nominal_datatype trm1 =
  Var1 "name"
| App1 "trm1" "trm1"
| Lam1 xs::"name list" t::"trm1" bind xs in t

thm trm1.strong_exhaust


text {* binding a finite set of names *}

nominal_datatype trm2 =
  Var2 "name"
| App2 "trm2" "trm2"
| Lam2 xs::"name fset" t::"trm2" bind (set) xs in t

thm trm2.strong_exhaust

text {* restricting a finite set of names *}

nominal_datatype trm3 =
  Var3 "name"
| App3 "trm3" "trm3"
| Lam3 xs::"name fset" t::"trm3" bind (res) xs in t

thm trm3.strong_exhaust

end