diff -r 4029105011ca -r 866208388c1d Nominal/LFex.thy --- a/Nominal/LFex.thy Thu Mar 11 17:47:29 2010 +0100 +++ b/Nominal/LFex.thy Thu Mar 11 17:49:07 2010 +0100 @@ -5,8 +5,6 @@ atom_decl name atom_decl ident -ML {* restricted_nominal := 2 *} - nominal_datatype kind = Type | KPi "ty" n::"name" k::"kind" bind n in k @@ -20,35 +18,6 @@ | App "trm" "trm" | Lam "ty" n::"name" t::"trm" bind n in t -lemma supports: - "{} supports Type" - "(supp (atom i)) supports (TConst i)" - "(supp A \ supp M) supports (TApp A M)" - "(supp (atom i)) supports (Const i)" - "(supp (atom x)) supports (Var x)" - "(supp M \ supp N) supports (App M N)" - "(supp ty \ supp (atom na) \ supp ki) supports (KPi ty na ki)" - "(supp ty \ supp (atom na) \ supp ty2) supports (TPi ty na ty2)" - "(supp ty \ supp (atom na) \ supp trm) supports (Lam ty na trm)" -apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh kind_ty_trm_perm) -apply(rule_tac [!] allI)+ -apply(rule_tac [!] impI) -apply(tactic {* ALLGOALS (REPEAT o etac conjE) *}) -apply(simp_all add: fresh_atom) -done - -lemma kind_ty_trm_fs: - "finite (supp (x\kind)) \ finite (supp (y\ty)) \ finite (supp (z\trm))" -apply(induct rule: kind_ty_trm_induct) -apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *}) -apply(simp_all add: supp_atom) -done - -instance kind and ty and trm :: fs -apply(default) -apply(simp_all only: kind_ty_trm_fs) -done - lemma ex_out: "(\x. Z x \ Q) = (Q \ (\x. Z x))" "(\x. Q \ Z x) = (Q \ (\x. Z x))"