| author | berghofe | 
| Mon, 12 Mar 2007 11:07:59 +0100 | |
| changeset 22436 | c9e384a956df | 
| parent 21627 | b822c7e61701 | 
| child 22933 | 338c7890c96f | 
| permissions | -rw-r--r-- | 
| 15803 | 1 | (* Title: Pure/Pure.thy | 
| 2 | ID: $Id$ | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 3 | *) | 
| 15803 | 4 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 5 | header {* The Pure theory *}
 | 
| 15803 | 6 | |
| 7 | theory Pure | |
| 8 | imports ProtoPure | |
| 9 | begin | |
| 19800 | 10 | |
| 19048 | 11 | setup  -- {* Common setup of internal components *}
 | 
| 15803 | 12 | |
| 20627 | 13 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 14 | subsection {* Meta-level connectives in assumptions *}
 | 
| 15803 | 15 | |
| 16 | lemma meta_mp: | |
| 18019 | 17 | assumes "PROP P ==> PROP Q" and "PROP P" | 
| 15803 | 18 | shows "PROP Q" | 
| 18019 | 19 | by (rule `PROP P ==> PROP Q` [OF `PROP P`]) | 
| 15803 | 20 | |
| 21 | lemma meta_spec: | |
| 18019 | 22 | assumes "!!x. PROP P(x)" | 
| 15803 | 23 | shows "PROP P(x)" | 
| 18019 | 24 | by (rule `!!x. PROP P(x)`) | 
| 15803 | 25 | |
| 26 | lemmas meta_allE = meta_spec [elim_format] | |
| 27 | ||
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 28 | |
| 21625 | 29 | subsection {* Embedded terms *}
 | 
| 30 | ||
| 31 | locale (open) meta_term_syntax = | |
| 32 |   fixes meta_term :: "'a => prop"  ("TERM _")
 | |
| 33 | ||
| 34 | parse_translation {*
 | |
| 21627 
b822c7e61701
meta_term_syntax: proper operation on untyped preterms;
 wenzelm parents: 
21625diff
changeset | 35 |   [("\<^fixed>meta_term", fn [t] => Const ("ProtoPure.term", dummyT --> propT) $ t)]
 | 
| 21625 | 36 | *} | 
| 37 | ||
| 38 | lemmas [intro?] = termI | |
| 39 | ||
| 40 | ||
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 41 | subsection {* Meta-level conjunction *}
 | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 42 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 43 | locale (open) meta_conjunction_syntax = | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 44 | fixes meta_conjunction :: "prop => prop => prop" (infixr "&&" 2) | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 45 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 46 | parse_translation {*
 | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 47 |   [("\<^fixed>meta_conjunction", fn [t, u] => Logic.mk_conjunction (t, u))]
 | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 48 | *} | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 49 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 50 | lemma all_conjunction: | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 51 | includes meta_conjunction_syntax | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 52 | shows "(!!x. PROP A(x) && PROP B(x)) == ((!!x. PROP A(x)) && (!!x. PROP B(x)))" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 53 | proof | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 54 | assume conj: "!!x. PROP A(x) && PROP B(x)" | 
| 19121 | 55 | show "(\<And>x. PROP A(x)) && (\<And>x. PROP B(x))" | 
| 56 | proof - | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 57 | fix x | 
| 19121 | 58 | from conj show "PROP A(x)" by (rule conjunctionD1) | 
| 59 | from conj show "PROP B(x)" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 60 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 61 | next | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 62 | assume conj: "(!!x. PROP A(x)) && (!!x. PROP B(x))" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 63 | fix x | 
| 19121 | 64 | show "PROP A(x) && PROP B(x)" | 
| 65 | proof - | |
| 66 | show "PROP A(x)" by (rule conj [THEN conjunctionD1, rule_format]) | |
| 67 | show "PROP B(x)" by (rule conj [THEN conjunctionD2, rule_format]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 68 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 69 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 70 | |
| 19121 | 71 | lemma imp_conjunction: | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 72 | includes meta_conjunction_syntax | 
| 19121 | 73 | shows "(PROP A ==> PROP B && PROP C) == (PROP A ==> PROP B) && (PROP A ==> PROP C)" | 
| 18836 | 74 | proof | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 75 | assume conj: "PROP A ==> PROP B && PROP C" | 
| 19121 | 76 | show "(PROP A ==> PROP B) && (PROP A ==> PROP C)" | 
| 77 | proof - | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 78 | assume "PROP A" | 
| 19121 | 79 | from conj [OF `PROP A`] show "PROP B" by (rule conjunctionD1) | 
| 80 | from conj [OF `PROP A`] show "PROP C" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 81 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 82 | next | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 83 | assume conj: "(PROP A ==> PROP B) && (PROP A ==> PROP C)" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 84 | assume "PROP A" | 
| 19121 | 85 | show "PROP B && PROP C" | 
| 86 | proof - | |
| 87 | from `PROP A` show "PROP B" by (rule conj [THEN conjunctionD1]) | |
| 88 | from `PROP A` show "PROP C" by (rule conj [THEN conjunctionD2]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 89 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 90 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 91 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 92 | lemma conjunction_imp: | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 93 | includes meta_conjunction_syntax | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 94 | shows "(PROP A && PROP B ==> PROP C) == (PROP A ==> PROP B ==> PROP C)" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 95 | proof | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 96 | assume r: "PROP A && PROP B ==> PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 97 | assume "PROP A" and "PROP B" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 98 | show "PROP C" by (rule r) - | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 99 | next | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 100 | assume r: "PROP A ==> PROP B ==> PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 101 | assume conj: "PROP A && PROP B" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 102 | show "PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 103 | proof (rule r) | 
| 19121 | 104 | from conj show "PROP A" by (rule conjunctionD1) | 
| 105 | from conj show "PROP B" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 106 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 107 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 108 | |
| 15803 | 109 | end |