author | ballarin |
Tue, 06 Jun 2006 10:05:57 +0200 | |
changeset 19783 | 82f365a14960 |
parent 19121 | d7fd5415a781 |
child 19800 | 5f764272183e |
permissions | -rw-r--r-- |
15803 | 1 |
(* Title: Pure/Pure.thy |
2 |
ID: $Id$ |
|
18710 | 3 |
|
19783 | 4 |
The Pure theory. |
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
5 |
*) |
15803 | 6 |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
7 |
header {* The Pure theory *} |
15803 | 8 |
|
9 |
theory Pure |
|
10 |
imports ProtoPure |
|
11 |
begin |
|
19783 | 12 |
ML{*set Toplevel.debug*} |
19048 | 13 |
setup -- {* Common setup of internal components *} |
15803 | 14 |
|
18710 | 15 |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
16 |
subsection {* Meta-level connectives in assumptions *} |
15803 | 17 |
|
18 |
lemma meta_mp: |
|
18019 | 19 |
assumes "PROP P ==> PROP Q" and "PROP P" |
15803 | 20 |
shows "PROP Q" |
18019 | 21 |
by (rule `PROP P ==> PROP Q` [OF `PROP P`]) |
15803 | 22 |
|
23 |
lemma meta_spec: |
|
18019 | 24 |
assumes "!!x. PROP P(x)" |
15803 | 25 |
shows "PROP P(x)" |
18019 | 26 |
by (rule `!!x. PROP P(x)`) |
15803 | 27 |
|
28 |
lemmas meta_allE = meta_spec [elim_format] |
|
29 |
||
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
30 |
|
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
31 |
subsection {* Meta-level conjunction *} |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
32 |
|
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
33 |
locale (open) meta_conjunction_syntax = |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
34 |
fixes meta_conjunction :: "prop => prop => prop" (infixr "&&" 2) |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
35 |
|
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
36 |
parse_translation {* |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
37 |
[("\<^fixed>meta_conjunction", fn [t, u] => Logic.mk_conjunction (t, u))] |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
38 |
*} |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
39 |
|
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
40 |
lemma all_conjunction: |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
41 |
includes meta_conjunction_syntax |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
42 |
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:
18019
diff
changeset
|
43 |
proof |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
44 |
assume conj: "!!x. PROP A(x) && PROP B(x)" |
19121 | 45 |
show "(\<And>x. PROP A(x)) && (\<And>x. PROP B(x))" |
46 |
proof - |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
47 |
fix x |
19121 | 48 |
from conj show "PROP A(x)" by (rule conjunctionD1) |
49 |
from conj show "PROP B(x)" by (rule conjunctionD2) |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
50 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
51 |
next |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
52 |
assume conj: "(!!x. PROP A(x)) && (!!x. PROP B(x))" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
53 |
fix x |
19121 | 54 |
show "PROP A(x) && PROP B(x)" |
55 |
proof - |
|
56 |
show "PROP A(x)" by (rule conj [THEN conjunctionD1, rule_format]) |
|
57 |
show "PROP B(x)" by (rule conj [THEN conjunctionD2, rule_format]) |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
58 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
59 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
60 |
|
19121 | 61 |
lemma imp_conjunction: |
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
62 |
includes meta_conjunction_syntax |
19121 | 63 |
shows "(PROP A ==> PROP B && PROP C) == (PROP A ==> PROP B) && (PROP A ==> PROP C)" |
18836 | 64 |
proof |
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
65 |
assume conj: "PROP A ==> PROP B && PROP C" |
19121 | 66 |
show "(PROP A ==> PROP B) && (PROP A ==> PROP C)" |
67 |
proof - |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
68 |
assume "PROP A" |
19121 | 69 |
from conj [OF `PROP A`] show "PROP B" by (rule conjunctionD1) |
70 |
from conj [OF `PROP A`] show "PROP C" by (rule conjunctionD2) |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
71 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
72 |
next |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
73 |
assume conj: "(PROP A ==> PROP B) && (PROP A ==> PROP C)" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
74 |
assume "PROP A" |
19121 | 75 |
show "PROP B && PROP C" |
76 |
proof - |
|
77 |
from `PROP A` show "PROP B" by (rule conj [THEN conjunctionD1]) |
|
78 |
from `PROP A` show "PROP C" by (rule conj [THEN conjunctionD2]) |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
79 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
80 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
81 |
|
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
82 |
lemma conjunction_imp: |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
83 |
includes meta_conjunction_syntax |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
84 |
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:
18019
diff
changeset
|
85 |
proof |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
86 |
assume r: "PROP A && PROP B ==> PROP C" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
87 |
assume "PROP A" and "PROP B" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
88 |
show "PROP C" by (rule r) - |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
89 |
next |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
90 |
assume r: "PROP A ==> PROP B ==> PROP C" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
91 |
assume conj: "PROP A && PROP B" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
92 |
show "PROP C" |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
93 |
proof (rule r) |
19121 | 94 |
from conj show "PROP A" by (rule conjunctionD1) |
95 |
from conj show "PROP B" by (rule conjunctionD2) |
|
18466
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
96 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
97 |
qed |
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
wenzelm
parents:
18019
diff
changeset
|
98 |
|
15803 | 99 |
end |