| 15803 |      1 | (*  Title:      Pure/Pure.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Larry Paulson and Makarius
 | 
|  |      4 | 
 | 
|  |      5 | The Pure theory.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | theory Pure
 | 
|  |      9 | imports ProtoPure
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | setup "Context.setup ()"
 | 
|  |     13 | 
 | 
|  |     14 | 
 | 
|  |     15 | text{*These meta-level elimination rules facilitate the use of assumptions
 | 
|  |     16 | that contain !! and ==>, especially in linear scripts. *}
 | 
|  |     17 | 
 | 
|  |     18 | lemma meta_mp:
 | 
|  |     19 |   assumes major: "PROP P ==> PROP Q" and minor: "PROP P"
 | 
|  |     20 |   shows "PROP Q"
 | 
| 15824 |     21 |     by (rule major [OF minor])
 | 
| 15803 |     22 | 
 | 
|  |     23 | lemma meta_spec:
 | 
|  |     24 |   assumes major: "!!x. PROP P(x)"
 | 
|  |     25 |   shows "PROP P(x)"
 | 
| 15824 |     26 |     by (rule major)
 | 
| 15803 |     27 | 
 | 
|  |     28 | lemmas meta_allE = meta_spec [elim_format]
 | 
|  |     29 | 
 | 
|  |     30 | end
 |