| author | haftmann | 
| Tue, 19 Sep 2006 15:22:03 +0200 | |
| changeset 20596 | 3950e65f48f8 | 
| parent 19800 | 5f764272183e | 
| child 20627 | 30da2841553e | 
| 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: 
18019 
diff
changeset
 | 
3  | 
*)  | 
| 15803 | 4  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
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  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
13  | 
subsection {* Meta-level connectives in assumptions *}
 | 
| 15803 | 14  | 
|
15  | 
lemma meta_mp:  | 
|
| 18019 | 16  | 
assumes "PROP P ==> PROP Q" and "PROP P"  | 
| 15803 | 17  | 
shows "PROP Q"  | 
| 18019 | 18  | 
by (rule `PROP P ==> PROP Q` [OF `PROP P`])  | 
| 15803 | 19  | 
|
20  | 
lemma meta_spec:  | 
|
| 18019 | 21  | 
assumes "!!x. PROP P(x)"  | 
| 15803 | 22  | 
shows "PROP P(x)"  | 
| 18019 | 23  | 
by (rule `!!x. PROP P(x)`)  | 
| 15803 | 24  | 
|
25  | 
lemmas meta_allE = meta_spec [elim_format]  | 
|
26  | 
||
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
27  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
28  | 
subsection {* Meta-level conjunction *}
 | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
29  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
30  | 
locale (open) meta_conjunction_syntax =  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
31  | 
fixes meta_conjunction :: "prop => prop => prop" (infixr "&&" 2)  | 
| 
 
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  | 
parse_translation {*
 | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
34  | 
  [("\<^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
 | 
35  | 
*}  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
36  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
37  | 
lemma all_conjunction:  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
38  | 
includes meta_conjunction_syntax  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
39  | 
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
 | 
40  | 
proof  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
41  | 
assume conj: "!!x. PROP A(x) && PROP B(x)"  | 
| 19121 | 42  | 
show "(\<And>x. PROP A(x)) && (\<And>x. PROP B(x))"  | 
43  | 
proof -  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
44  | 
fix x  | 
| 19121 | 45  | 
from conj show "PROP A(x)" by (rule conjunctionD1)  | 
46  | 
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
 | 
47  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
48  | 
next  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
49  | 
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
 | 
50  | 
fix x  | 
| 19121 | 51  | 
show "PROP A(x) && PROP B(x)"  | 
52  | 
proof -  | 
|
53  | 
show "PROP A(x)" by (rule conj [THEN conjunctionD1, rule_format])  | 
|
54  | 
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
 | 
55  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
56  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
57  | 
|
| 19121 | 58  | 
lemma imp_conjunction:  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
59  | 
includes meta_conjunction_syntax  | 
| 19121 | 60  | 
shows "(PROP A ==> PROP B && PROP C) == (PROP A ==> PROP B) && (PROP A ==> PROP C)"  | 
| 18836 | 61  | 
proof  | 
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
62  | 
assume conj: "PROP A ==> PROP B && PROP C"  | 
| 19121 | 63  | 
show "(PROP A ==> PROP B) && (PROP A ==> PROP C)"  | 
64  | 
proof -  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
65  | 
assume "PROP A"  | 
| 19121 | 66  | 
from conj [OF `PROP A`] show "PROP B" by (rule conjunctionD1)  | 
67  | 
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
 | 
68  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
69  | 
next  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
70  | 
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
 | 
71  | 
assume "PROP A"  | 
| 19121 | 72  | 
show "PROP B && PROP C"  | 
73  | 
proof -  | 
|
74  | 
from `PROP A` show "PROP B" by (rule conj [THEN conjunctionD1])  | 
|
75  | 
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
 | 
76  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
77  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
78  | 
|
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
79  | 
lemma conjunction_imp:  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
80  | 
includes meta_conjunction_syntax  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
81  | 
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
 | 
82  | 
proof  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
83  | 
assume r: "PROP A && PROP B ==> PROP C"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
84  | 
assume "PROP A" and "PROP B"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
85  | 
show "PROP C" by (rule r) -  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
86  | 
next  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
87  | 
assume r: "PROP A ==> PROP B ==> PROP C"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
88  | 
assume conj: "PROP A && PROP B"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
89  | 
show "PROP C"  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
90  | 
proof (rule r)  | 
| 19121 | 91  | 
from conj show "PROP A" by (rule conjunctionD1)  | 
92  | 
from conj show "PROP B" by (rule conjunctionD2)  | 
|
| 
18466
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
93  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
94  | 
qed  | 
| 
 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 
wenzelm 
parents: 
18019 
diff
changeset
 | 
95  | 
|
| 15803 | 96  | 
end  |