| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 14 Jan 2009 23:44:14 +0000 | |
| changeset 72 | 7b8c4fe235aa | 
| parent 69 | 19106a9975c1 | 
| child 74 | f6f8f8ba1eb1 | 
| permissions | -rw-r--r-- | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1  | 
theory NamedThms  | 
| 58 | 2  | 
imports "../Base"  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
4  | 
|
| 
13
 
2b07da8b310d
polished and added a subdirectory for the recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
12 
diff
changeset
 | 
5  | 
section {* Accumulate a List of Theorems under a Name *} 
 | 
| 
2
 
978a3c2ed7ce
split the document into smaller pieces;
 
Christian Urban <urbanc@in.tum.de> 
parents: 
0 
diff
changeset
 | 
6  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
8  | 
text {*
 | 
| 20 | 9  | 
  {\bf Problem:} 
 | 
10  | 
  Your tool @{text foo} works with special rules, called @{text foo}-rules.
 | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
11  | 
  Users should be able to declare @{text foo}-rules in the theory,
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
12  | 
which are then used in a method.\smallskip  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
13  | 
|
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
14  | 
  {\bf Solution:} This can be achieved using named theorem lists.\smallskip
 | 
| 
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
15  | 
|
| 
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
16  | 
Named theorem lists can be set up using the code  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
18  | 
*}  | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
19  | 
|
| 
69
 
19106a9975c1
highligted the background of ML-code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
68 
diff
changeset
 | 
20  | 
ML{*structure FooRules = NamedThmsFun (
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
21  | 
val name = "foo"  | 
| 
69
 
19106a9975c1
highligted the background of ML-code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
68 
diff
changeset
 | 
22  | 
val description = "Rules for foo"); *}  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
24  | 
text {*
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
25  | 
and the command  | 
| 
68
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
63 
diff
changeset
 | 
26  | 
*}  | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
27  | 
|
| 
68
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
63 
diff
changeset
 | 
28  | 
setup {* FooRules.setup *}
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
29  | 
|
| 
68
 
e7519207c2b7
added more to the "new command section" and tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
63 
diff
changeset
 | 
30  | 
text {*
 | 
| 20 | 31  | 
This code declares a context data slot where the theorems are stored,  | 
| 58 | 32  | 
  an attribute @{text foo} (with the usual @{text add} and @{text del} options
 | 
| 
51
 
c346c156a7cd
completes the recipie on antiquotations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
48 
diff
changeset
 | 
33  | 
for adding and deleting theorems) and an internal ML interface to retrieve and  | 
| 20 | 34  | 
modify the theorems.  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
|
| 20 | 36  | 
Furthermore, the facts are made available on the user level under the dynamic  | 
| 
51
 
c346c156a7cd
completes the recipie on antiquotations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
48 
diff
changeset
 | 
37  | 
  fact name @{text foo}. For example we can declare three lemmas to be of the kind
 | 
| 58 | 38  | 
  @{text foo} by:
 | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
39  | 
*}  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
41  | 
lemma rule1[foo]: "A" sorry  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
42  | 
lemma rule2[foo]: "B" sorry  | 
| 
48
 
609f9ef73494
fixed FIXME's in fake responses
 
Christian Urban <urbanc@in.tum.de> 
parents: 
47 
diff
changeset
 | 
43  | 
lemma rule3[foo]: "C" sorry  | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
44  | 
|
| 
51
 
c346c156a7cd
completes the recipie on antiquotations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
48 
diff
changeset
 | 
45  | 
text {* and undeclare the first one by: *}
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
46  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
47  | 
declare rule1[foo del]  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
48  | 
|
| 63 | 49  | 
text {* and query the remaining ones with:
 | 
| 
47
 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 
Christian Urban <urbanc@in.tum.de> 
parents: 
43 
diff
changeset
 | 
50  | 
|
| 
72
 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 
Christian Urban <urbanc@in.tum.de> 
parents: 
69 
diff
changeset
 | 
51  | 
@{ML_response_fake_both [display,gray] "thm foo" 
 | 
| 63 | 52  | 
"?C  | 
53  | 
?B"}  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
54  | 
|
| 58 | 55  | 
  On the ML-level the rules marked with @{text "foo"} an be retrieved
 | 
| 
51
 
c346c156a7cd
completes the recipie on antiquotations
 
Christian Urban <urbanc@in.tum.de> 
parents: 
48 
diff
changeset
 | 
56  | 
  using the function @{ML FooRules.get}:
 | 
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
57  | 
|
| 
72
 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 
Christian Urban <urbanc@in.tum.de> 
parents: 
69 
diff
changeset
 | 
58  | 
  @{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"}
 | 
| 20 | 59  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
60  | 
  \begin{readmore}
 | 
| 63 | 61  | 
  For more information see @{ML_file "Pure/Tools/named_thms.ML"} and also
 | 
62  | 
  the recipe in Section~\ref{recipe:storingdata} about storing arbitrary
 | 
|
63  | 
data.  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
64  | 
  \end{readmore}
 | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
65  | 
*}  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
66  | 
|
| 
15
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
67  | 
text {*
 | 
| 
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
68  | 
(FIXME: maybe add a comment about the case when the theorems  | 
| 
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
69  | 
to be added need to satisfy certain properties)  | 
| 
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
70  | 
|
| 
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
71  | 
*}  | 
| 
 
9da9ba2b095b
added a solution section and some other minor additions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
14 
diff
changeset
 | 
72  | 
|
| 
0
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
74  | 
end  | 
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
02503850a8cf
initial commit of Alexander's files
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
78  |