CookBook/Recipes/NamedThms.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 26 Feb 2009 14:20:52 +0000
changeset 150 cb39c41548bd
parent 119 4536782969fa
permissions -rw-r--r--
added a comment about Conv.fun_conv
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory NamedThms
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
     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
119
4536782969fa added an acknowledgement section
Christian Urban <urbanc@in.tum.de>
parents: 74
diff changeset
     5
section {* Accumulate a List of Theorems under a Name\label{rec:named} *} 
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
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
     9
  {\bf Problem:} 
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
    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
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
    31
  This code declares a context data slot where the theorems are stored,
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    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
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
    34
  modify the theorems.
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
20
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
    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
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    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
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    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
74
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    51
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    52
  \isacommand{thm}~@{text "foo"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    53
  @{text "> ?C"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    54
  @{text "> ?B"}\\
Christian Urban <urbanc@in.tum.de>
parents: 72
diff changeset
    55
  \end{isabelle}
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    57
  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
    58
  using the function @{ML FooRules.get}:
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
    60
  @{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"}
20
5ae6a1bb91c9 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 15
diff changeset
    61
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  \begin{readmore}
63
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    63
  For more information see @{ML_file "Pure/Tools/named_thms.ML"} and also
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    64
  the recipe in Section~\ref{recipe:storingdata} about storing arbitrary
Christian Urban <urbanc@in.tum.de>
parents: 58
diff changeset
    65
  data.
0
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  \end{readmore}
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
 *}
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
    69
text {*
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
    70
  (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
    71
  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
    72
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
    73
*}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
    74
0
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
end
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
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
02503850a8cf initial commit of Alexander's files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80