src/Pure/Pure.thy
author wenzelm
Mon, 04 Jul 2005 17:07:13 +0200
changeset 16679 abd1461fa288
parent 15824 222eeb9655f3
child 18019 d1ff9ebb8bcb
permissions -rw-r--r--
dest_ctyp: raise exception for non-constructor; dest_comb: replaced expensive fastype_of by Term.argument_type; dest_abs: replaced expensive variant_abs by Term.dest_abs; hyps: fast_term_ord;

(*  Title:      Pure/Pure.thy
    ID:         $Id$
    Author:     Larry Paulson and Makarius

The Pure theory.
*)

theory Pure
imports ProtoPure
begin

setup "Context.setup ()"


text{*These meta-level elimination rules facilitate the use of assumptions
that contain !! and ==>, especially in linear scripts. *}

lemma meta_mp:
  assumes major: "PROP P ==> PROP Q" and minor: "PROP P"
  shows "PROP Q"
    by (rule major [OF minor])

lemma meta_spec:
  assumes major: "!!x. PROP P(x)"
  shows "PROP P(x)"
    by (rule major)

lemmas meta_allE = meta_spec [elim_format]

end