author | Christian Urban <urbanc@in.tum.de> |
Mon, 27 Aug 2012 10:24:10 +0100 | |
changeset 535 | 5734ab5dd86d |
parent 529 | 13d7ea419c5f |
child 562 | daf404920ab9 |
permissions | -rw-r--r-- |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
theory Introspection |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
imports "../Appendix" |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
section {* Introspection of Theorems and Proofs \label{rec:introspection} *} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
text{* |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
{\bf Problem:} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
How to obtain all theorems that are used in the proof of a theorem?\smallskip |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
{\bf Solution:} They can be obtained by introspecting the theorem.\smallskip |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
To introspect a theorem, let us define the following three functions that |
528 | 15 |
analyse the @{ML_type_ind proof_body} data-structure from the structure |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
16 |
@{ML_struct Proofterm}. |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
*} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
ML %grayML{*fun pthms_of (PBody {thms, ...}) = map #2 thms |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
val get_names = map #1 o pthms_of |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
val get_pbodies = map (Future.join o #3) o pthms_of *} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
text {* |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
24 |
To see what their purpose is, consider the following three short proofs. |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
*} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
lemma my_conjIa: |
528 | 28 |
shows "A \<and> B \<Longrightarrow> A \<and> B" |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
apply(rule conjI) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
apply(drule conjunct1) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
apply(assumption) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
apply(drule conjunct2) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
apply(assumption) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
done |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
lemma my_conjIb: |
528 | 37 |
shows "A \<and> B \<Longrightarrow> A \<and> B" |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
apply(assumption) |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
done |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
|
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
41 |
lemma my_conjIc: |
528 | 42 |
shows "A \<and> B \<Longrightarrow> A \<and> B" |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
43 |
apply(auto) |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
44 |
done |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
45 |
|
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
text {* |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
47 |
While the information about which theorems are used is obvious in |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
48 |
the first two proofs, it is not obvious in the third, because of the |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
49 |
@{text auto}-step. Fortunately, ``behind'' every proved theorem is |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
50 |
a proof-tree that records all theorems that are employed for |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
51 |
establishing this theorem. We can traverse this proof-tree |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
52 |
extracting this information. Let us first extract the name of the |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
53 |
established theorem. This can be done with |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
@{ML_response [display,gray] |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
"@{thm my_conjIa} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
|> Thm.proof_body_of |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
|> get_names" |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
"[\"Introspection.my_conjIa\"]"} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
|
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
61 |
whereby @{text "Introspection"} refers to the theory name in which |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
62 |
we established the theorem @{thm [source] my_conjIa}. The function @{ML_ind |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
63 |
proof_body_of in Thm} returns a part of the data that is stored in a |
529
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
64 |
theorem. Notice that the first proof above references |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
65 |
three theorems, namely @{thm [source] conjI}, @{thm [source] conjunct1} |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
66 |
and @{thm [source] conjunct2}. We can obtain them by descending into the |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
67 |
first level of the proof-tree, as follows. |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
@{ML_response [display,gray] |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
"@{thm my_conjIa} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
|> Thm.proof_body_of |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
|> get_pbodies |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
|> map get_names |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
|> List.concat" |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
"[\"HOL.conjunct2\", \"HOL.conjunct1\", \"HOL.conjI\", \"Pure.protectD\", |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
\"Pure.protectI\"]"} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
|
529
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
78 |
The theorems @{thm [source] protectD} and @{thm [source] |
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
79 |
protectI} that are internal theorems are always part of a |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
proof in Isabelle. Note also that applications of @{text assumption} do not |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
count as a separate theorem, as you can see in the following code. |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
@{ML_response [display,gray] |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
"@{thm my_conjIb} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
85 |
|> Thm.proof_body_of |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
86 |
|> get_pbodies |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
87 |
|> map get_names |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
|> List.concat" |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
"[\"Pure.protectD\", \"Pure.protectI\"]"} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
90 |
|
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
91 |
Interestingly, but not surprisingly, the proof of @{thm [source] |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
92 |
my_conjIc} procceeds quite differently from @{thm [source] my_conjIa} |
529
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
93 |
and @{thm [source] my_conjIb}, as can be seen by the theorems that |
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
94 |
are returned for @{thm [source] my_conjIc}. |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
95 |
|
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
96 |
@{ML_response [display,gray] |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
97 |
"@{thm my_conjIc} |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
98 |
|> Thm.proof_body_of |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
99 |
|> get_pbodies |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
100 |
|> map get_names |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
101 |
|> List.concat" |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
102 |
"[\"HOL.Eq_TrueI\", \"HOL.simp_thms_25\", \"HOL.eq_reflection\", |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
103 |
\"HOL.conjunct2\", \"HOL.conjunct1\", \"HOL.TrueI\", \"Pure.protectD\", |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
104 |
\"Pure.protectI\"]"} |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
105 |
|
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
106 |
Of course we can also descend into the second level of the tree |
529
13d7ea419c5f
moved the introspection part into the theorem section
Christian Urban <urbanc@in.tum.de>
parents:
528
diff
changeset
|
107 |
``behind'' @{thm [source] my_conjIa} say, which |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
108 |
means we obtain the theorems that are used in order to prove |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
109 |
@{thm [source] conjunct1}, @{thm [source] conjunct2} and @{thm [source] conjI}. |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
110 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
111 |
@{ML_response [display, gray] |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
112 |
"@{thm my_conjIa} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
113 |
|> Thm.proof_body_of |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
114 |
|> get_pbodies |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
115 |
|> map get_pbodies |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
116 |
|> (map o map) get_names |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
|> List.concat |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
|> List.concat |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
|> duplicates (op=)" |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
120 |
"[\"HOL.spec\", \"HOL.and_def\", \"HOL.mp\", \"HOL.impI\", \"Pure.protectD\", |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
\"Pure.protectI\"]"} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
123 |
\begin{readmore} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
The data-structure @{ML_type proof_body} is implemented |
527
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
125 |
in the file @{ML_file "Pure/proofterm.ML"}. The implementation |
efe63c062e48
polished introspection section
Christian Urban <urbanc@in.tum.de>
parents:
525
diff
changeset
|
126 |
of theorems and related functions are in @{ML_file "Pure/thm.ML"}. |
525
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
127 |
\end{readmore} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
*} |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
132 |
end |
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
|
92a3600e50e4
added a new recipe for introspecting theorems (suggested by Lukas and Rafal)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |