39 |
42 |
40 Like normal Isabelle proof scripts, \isacommand{ML}-commands can be |
43 Like normal Isabelle proof scripts, \isacommand{ML}-commands can be |
41 evaluated by using the advance and undo buttons of your Isabelle |
44 evaluated by using the advance and undo buttons of your Isabelle |
42 environment. The code inside the \isacommand{ML}-command can also contain |
45 environment. The code inside the \isacommand{ML}-command can also contain |
43 value and function bindings, and even those can be undone when the proof |
46 value and function bindings, and even those can be undone when the proof |
44 script is retracted. As mentioned earlier, we will drop the |
47 script is retracted. As mentioned in the Introduction, we will drop the |
45 \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we |
48 \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we |
46 show code. The lines prefixed with @{text [quotes] ">"} are not part of the |
49 show code. The lines prefixed with @{text [quotes] ">"} are not part of the |
47 code, rather they indicate what the response is when the code is evaluated. |
50 code, rather they indicate what the response is when the code is evaluated. |
48 |
51 |
49 Once a portion of code is relatively stable, you usually want to export it |
52 Once a portion of code is relatively stable, you usually want to export it |
149 ML{*fun str_of_cterms ctxt ts = |
152 ML{*fun str_of_cterms ctxt ts = |
150 commas (map (str_of_cterm ctxt) ts)*} |
153 commas (map (str_of_cterm ctxt) ts)*} |
151 |
154 |
152 text {* |
155 text {* |
153 The easiest way to get the string of a theorem is to transform it |
156 The easiest way to get the string of a theorem is to transform it |
154 into a @{ML_type cterm} using the function @{ML crep_thm}. |
157 into a @{ML_type cterm} using the function @{ML crep_thm}. Theorems |
155 *} |
158 also include schematic variables, such as @{text "?P"}, @{text "?Q"} |
156 |
159 and so on. In order to improve the readability of theorems we convert |
157 ML{*fun str_of_thm ctxt thm = |
160 these schematic variables into free variables using the |
158 let |
161 function @{ML Variable.import_thms}. |
159 val {prop, ...} = crep_thm thm |
162 *} |
160 in |
163 |
161 str_of_cterm ctxt prop |
164 ML{*fun no_vars ctxt thm = |
162 end*} |
165 let |
|
166 val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt |
|
167 in |
|
168 thm' |
|
169 end |
|
170 |
|
171 fun str_of_thm ctxt thm = |
|
172 let |
|
173 val {prop, ...} = crep_thm (no_vars ctxt thm) |
|
174 in |
|
175 str_of_cterm ctxt prop |
|
176 end*} |
163 |
177 |
164 text {* |
178 text {* |
165 Again the function @{ML commas} helps with printing more than one theorem. |
179 Again the function @{ML commas} helps with printing more than one theorem. |
166 *} |
180 *} |
167 |
181 |
168 ML{*fun str_of_thms ctxt thms = |
182 ML{*fun str_of_thms ctxt thms = |
169 commas (map (str_of_thm ctxt) thms)*} |
183 commas (map (str_of_thm ctxt) thms)*} |
170 |
|
171 |
184 |
172 section {* Combinators\label{sec:combinators} *} |
185 section {* Combinators\label{sec:combinators} *} |
173 |
186 |
174 text {* |
187 text {* |
175 For beginners perhaps the most puzzling parts in the existing code of Isabelle are |
188 For beginners perhaps the most puzzling parts in the existing code of Isabelle are |
393 function that extracts the theorem names stored in a simpset. |
406 function that extracts the theorem names stored in a simpset. |
394 *} |
407 *} |
395 |
408 |
396 ML{*fun get_thm_names_from_ss simpset = |
409 ML{*fun get_thm_names_from_ss simpset = |
397 let |
410 let |
398 val ({rules,...}, _) = MetaSimplifier.rep_ss simpset |
411 val {simps,...} = MetaSimplifier.dest_ss simpset |
399 in |
412 in |
400 map #name (Net.entries rules) |
413 map #1 simps |
401 end*} |
414 end*} |
402 |
415 |
403 text {* |
416 text {* |
404 The function @{ML rep_ss in MetaSimplifier} returns a record containing all |
417 The function @{ML dest_ss in MetaSimplifier} returns a record containing all |
405 information about the simpset. The rules of a simpset are stored in a |
418 information stored in the simpset. We are only interested in the You can now use @{ML get_thm_names_from_ss} to obtain all names of theorems stored |
406 \emph{discrimination net} (a data structure for fast indexing). From this |
419 in the current simpset. This simpset can be referred to using the antiquotation |
407 net you can extract the entries using the function @{ML Net.entries}. |
|
408 You can now use @{ML get_thm_names_from_ss} to obtain all names of theorems stored |
|
409 in the current simpset---this simpset can be referred to using the antiquotation |
|
410 @{text "@{simpset}"}. |
420 @{text "@{simpset}"}. |
411 |
421 |
412 @{ML_response_fake [display,gray] |
422 @{ML_response_fake [display,gray] |
413 "get_thm_names_from_ss @{simpset}" |
423 "get_thm_names_from_ss @{simpset}" |
414 "[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"} |
424 "[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"} |
415 |
425 |
416 \begin{readmore} |
426 Again, this way or referencing simpsets makes you independent from additions |
417 The infrastructure for simpsets is implemented in @{ML_file "Pure/meta_simplifier.ML"} |
427 of lemmas to the simpset by the user that potentially cause loops. |
418 and @{ML_file "Pure/simplifier.ML"}. Discrimination nets are implemented |
|
419 in @{ML_file "Pure/net.ML"}. |
|
420 \end{readmore} |
|
421 |
428 |
422 While antiquotations have many applications, they were originally introduced in order |
429 While antiquotations have many applications, they were originally introduced in order |
423 to avoid explicit bindings for theorems such as: |
430 to avoid explicit bindings for theorems such as: |
424 *} |
431 *} |
425 |
432 |
426 ML{*val allI = thm "allI" *} |
433 ML{*val allI = thm "allI" *} |
427 |
434 |
428 text {* |
435 text {* |
429 These bindings are difficult to maintain and also can be accidentally |
436 These bindings are difficult to maintain and also can be accidentally |
430 overwritten by the user. This often breakes Isabelle |
437 overwritten by the user. This often broke Isabelle |
431 packages. Antiquotations solve this problem, since they are ``linked'' |
438 packages. Antiquotations solve this problem, since they are ``linked'' |
432 statically at compile-time. However, this static linkage also limits their |
439 statically at compile-time. However, this static linkage also limits their |
433 usefulness in cases where data needs to be build up dynamically. In the |
440 usefulness in cases where data needs to be build up dynamically. In the |
434 course of this chapter you will learn more about these antiquotations: |
441 course of this chapter you will learn more about these antiquotations: |
435 they can simplify Isabelle programming since one can directly access all |
442 they can simplify Isabelle programming since one can directly access all |
436 kinds of logical elements from th ML-level. |
443 kinds of logical elements from th ML-level. |
437 |
|
438 *} |
444 *} |
439 |
445 |
440 section {* Terms and Types *} |
446 section {* Terms and Types *} |
441 |
447 |
442 text {* |
448 text {* |
543 To see this apply @{text "@{term S}"}, @{text "@{term T}"} and @{text "@{typ nat}"} |
549 To see this apply @{text "@{term S}"}, @{text "@{term T}"} and @{text "@{typ nat}"} |
544 to both functions. With @{ML make_imp} we obtain the intended term involving |
550 to both functions. With @{ML make_imp} we obtain the intended term involving |
545 the given arguments |
551 the given arguments |
546 |
552 |
547 @{ML_response [display,gray] "make_imp @{term S} @{term T} @{typ nat}" |
553 @{ML_response [display,gray] "make_imp @{term S} @{term T} @{typ nat}" |
548 "Const \<dots> $ |
554 "Const \<dots> $ |
549 Abs (\"x\", Type (\"nat\",[]), |
555 Abs (\"x\", Type (\"nat\",[]), |
550 Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ |
556 Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ (Free (\"T\",\<dots>) $ \<dots>))"} |
551 (Free (\"T\",\<dots>) $ \<dots>))"} |
|
552 |
557 |
553 whereas with @{ML make_wrong_imp} we obtain a term involving the @{term "P"} |
558 whereas with @{ML make_wrong_imp} we obtain a term involving the @{term "P"} |
554 and @{text "Q"} from the antiquotation. |
559 and @{text "Q"} from the antiquotation. |
555 |
560 |
556 @{ML_response [display,gray] "make_wrong_imp @{term S} @{term T} @{typ nat}" |
561 @{ML_response [display,gray] "make_wrong_imp @{term S} @{term T} @{typ nat}" |
557 "Const \<dots> $ |
562 "Const \<dots> $ |
558 Abs (\"x\", \<dots>, |
563 Abs (\"x\", \<dots>, |
559 Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ |
564 Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ |
560 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"} |
565 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"} |
561 |
566 |
562 Although types of terms can often be inferred, there are many |
567 Although types of terms can often be inferred, there are many |
563 situations where you need to construct types manually, especially |
568 situations where you need to construct types manually, especially |
564 when defining constants. For example the function returning a function |
569 when defining constants. For example the function returning a function |
565 type is as follows: |
570 type is as follows: |
566 |
571 |
567 *} |
572 *} |
568 |
573 |
569 ML{*fun make_fun_type tau1 tau2 = Type ("fun",[tau1,tau2]) *} |
574 ML{*fun make_fun_type tau1 tau2 = Type ("fun", [tau1, tau2]) *} |
570 |
575 |
571 text {* This can be equally written as: *} |
576 text {* This can be equally written with the combinator @{ML "-->"} as: *} |
572 |
577 |
573 ML{*fun make_fun_type tau1 tau2 = tau1 --> tau2 *} |
578 ML{*fun make_fun_type tau1 tau2 = tau1 --> tau2 *} |
574 |
579 |
575 text {* |
580 text {* |
576 A handy function for manipulating terms is @{ML map_types}: it takes a |
581 A handy function for manipulating terms is @{ML map_types}: it takes a |
577 function and applies it to every type in the term. You can, for example, |
582 function and applies it to every type in a term. You can, for example, |
578 change every @{typ nat} in a term into an @{typ int} using the function: |
583 change every @{typ nat} in a term into an @{typ int} using the function: |
579 *} |
584 *} |
580 |
585 |
581 ML{*fun nat_to_int t = |
586 ML{*fun nat_to_int t = |
582 (case t of |
587 (case t of |
699 | is_nil_or_all _ = false *} |
704 | is_nil_or_all _ = false *} |
700 |
705 |
701 text {* |
706 text {* |
702 Occasional you have to calculate what the ``base'' name of a given |
707 Occasional you have to calculate what the ``base'' name of a given |
703 constant is. For this you can use the function @{ML Sign.extern_const} or |
708 constant is. For this you can use the function @{ML Sign.extern_const} or |
704 @{ML Sign.base_name}. For example: |
709 @{ML Long_Name.base_name}. For example: |
705 |
710 |
706 @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""} |
711 @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""} |
707 |
712 |
708 The difference between both functions is that @{ML extern_const in Sign} returns |
713 The difference between both functions is that @{ML extern_const in Sign} returns |
709 the smallest name that is still unique, whereas @{ML base_name in Sign} always |
714 the smallest name that is still unique, whereas @{ML base_name in Long_Name} always |
710 strips off all qualifiers. |
715 strips off all qualifiers. |
711 |
716 |
712 \begin{readmore} |
717 \begin{readmore} |
713 FIXME |
718 Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"}; |
|
719 functions about signatures in @{ML_file "Pure/sign.ML"}. |
|
720 |
714 \end{readmore} |
721 \end{readmore} |
715 *} |
722 *} |
716 |
723 |
717 |
724 |
718 section {* Type-Checking *} |
725 section {* Type-Checking *} |
892 "[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags |
916 "[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags |
893 annotated to theorems, but functions that do further processing once a |
917 annotated to theorems, but functions that do further processing once a |
894 theorem is proven. In particular, it is not possible to find out |
918 theorem is proven. In particular, it is not possible to find out |
895 what are all theorems that have a given attribute in common, unless of course |
919 what are all theorems that have a given attribute in common, unless of course |
896 the function behind the attribute stores the theorems in a retrivable |
920 the function behind the attribute stores the theorems in a retrivable |
897 datastructure. This can be easily done by the recipe described in |
921 datastructure. |
898 \ref{rec:named}. |
|
899 |
922 |
900 If you want to print out all currently known attributes a theorem |
923 If you want to print out all currently known attributes a theorem |
901 can have, you can use the function: |
924 can have, you can use the function: |
902 |
925 |
903 @{ML_response_fake [display,gray] "Attrib.print_attributes @{theory}" |
926 @{ML_response_fake [display,gray] "Attrib.print_attributes @{theory}" |
917 text {* |
940 text {* |
918 where the function @{ML "Thm.rule_attribute"} expects a function taking a |
941 where the function @{ML "Thm.rule_attribute"} expects a function taking a |
919 context (which we ignore in the code above) and a theorem (@{text thm}), and |
942 context (which we ignore in the code above) and a theorem (@{text thm}), and |
920 returns another theorem (namely @{text thm} resolved with the lemma |
943 returns another theorem (namely @{text thm} resolved with the lemma |
921 @{thm [source] sym}: @{thm sym[no_vars]}). The function @{ML "Thm.rule_attribute"} then returns |
944 @{thm [source] sym}: @{thm sym[no_vars]}). The function @{ML "Thm.rule_attribute"} then returns |
922 an attribute (of type @{ML_type "attribute"}). |
945 an attribute. |
923 |
946 |
924 Before we can use the attribute, we need to set it up. This can be done |
947 Before we can use the attribute, we need to set it up. This can be done |
925 using the function @{ML Attrib.add_attributes} as follows. |
948 using the function @{ML Attrib.add_attributes} as follows. |
926 *} |
949 *} |
927 |
950 |
928 setup {* |
951 setup {* |
929 Attrib.add_attributes |
952 Attrib.add_attributes |
930 [("my_sym", Attrib.no_args my_symmetric, "applying the sym rule")] |
953 [("my_sym", Attrib.no_args my_symmetric, "applying the sym rule")] *} |
931 *} |
|
932 |
954 |
933 text {* |
955 text {* |
934 The attribute does not expect any further arguments (unlike @{text "[OF |
956 The attribute does not expect any further arguments (unlike @{text "[OF |
935 \<dots>]"}, for example, which can take a list of theorems as argument). Therefore |
957 \<dots>]"}, for example, which can take a list of theorems as argument). Therefore |
936 we use the function @{ML Attrib.no_args}. Later on we will also consider |
958 we use the function @{ML Attrib.no_args}. Later on we will also consider |
947 \begin{isabelle} |
969 \begin{isabelle} |
948 \isacommand{thm}~@{text "test[my_sym]"}\\ |
970 \isacommand{thm}~@{text "test[my_sym]"}\\ |
949 @{text "> "}~@{thm test[my_sym]} |
971 @{text "> "}~@{thm test[my_sym]} |
950 \end{isabelle} |
972 \end{isabelle} |
951 |
973 |
952 *} |
974 The purpose of @{ML Thm.rule_attribute} is to directly manipulate theorems. |
953 |
975 Another usage of attributes is to add and delete theorems from stored data. |
954 text {* |
976 For example the attribute @{text "[simp]"} adds or deletes a theorem from the |
955 What are: |
977 current simpset. For these applications, you can use @{ML Thm.declaration_attribute}. |
956 |
978 To illustrate this function, let us introduce a reference containing a list |
957 |
979 of theorems. |
958 @{text "declaration_attribute"} |
980 *} |
959 |
981 |
960 @{text "theory_attributes"} |
982 ML{*val my_thms = ref ([]:thm list)*} |
961 |
983 |
962 @{text "proof_attributes"} |
984 text {* |
|
985 A word of warning: such references must not be used in any code that |
|
986 is meant to be more than just for testing purposes! Here it is only used |
|
987 to illustrate matters. We will show later how to store data properly without |
|
988 using references. The function @{ML Thm.declaration_attribute} expects us to |
|
989 provide two functions that add and delete theorems from this list. |
|
990 For this we use the two functions: |
|
991 *} |
|
992 |
|
993 ML{*fun my_thms_add thm ctxt = |
|
994 (my_thms := Thm.add_thm thm (!my_thms); ctxt) |
|
995 |
|
996 fun my_thms_del thm ctxt = |
|
997 (my_thms := Thm.del_thm thm (!my_thms); ctxt)*} |
|
998 |
|
999 text {* |
|
1000 These functions take a theorem and a context and, for what we are explaining |
|
1001 here it is sufficient that they just return the context unchanged. They change |
|
1002 however the reference @{ML my_thms}, whereby the function @{ML Thm.add_thm} |
|
1003 adds a theorem if it is not already included in the list, and @{ML |
|
1004 Thm.del_thm} deletes one. Both functions use the predicate @{ML |
|
1005 Thm.eq_thm_prop} which compares theorems according to their proved |
|
1006 propositions (modulo alpha-equivalence). |
|
1007 |
|
1008 |
|
1009 You can turn both functions into attributes using |
|
1010 *} |
|
1011 |
|
1012 ML{*val my_add = Thm.declaration_attribute my_thms_add |
|
1013 val my_del = Thm.declaration_attribute my_thms_del *} |
|
1014 |
|
1015 text {* |
|
1016 and set up the attributes as follows |
|
1017 *} |
|
1018 |
|
1019 setup {* |
|
1020 Attrib.add_attributes |
|
1021 [("my_thms", Attrib.add_del_args my_add my_del, |
|
1022 "maintaining a list of my_thms")] *} |
|
1023 |
|
1024 text {* |
|
1025 Now if you prove the lemma attaching the attribute @{text "[my_thms]"} |
|
1026 *} |
|
1027 |
|
1028 lemma trueI_2[my_thms]: "True" by simp |
|
1029 |
|
1030 text {* |
|
1031 then you can see the lemma is added to the initially empty list. |
|
1032 |
|
1033 @{ML_response_fake [display,gray] |
|
1034 "!my_thms" "[\"True\"]"} |
|
1035 |
|
1036 You can also add theorems using the command \isacommand{declare}. |
|
1037 *} |
|
1038 |
|
1039 declare test[my_thms] trueI_2[my_thms add] |
|
1040 |
|
1041 text {* |
|
1042 The @{text "add"} is the default operation and does not need |
|
1043 to be given. This declaration will cause the theorem list to be |
|
1044 updated as follows. |
|
1045 |
|
1046 @{ML_response_fake [display,gray] |
|
1047 "!my_thms" |
|
1048 "[\"True\", \"Suc (Suc 0) = 2\"]"} |
|
1049 |
|
1050 The theorem @{thm [source] trueI_2} only appears once, since the |
|
1051 function @{ML Thm.add_thm} tests for duplicates, before extending |
|
1052 the list. Deletion from the list works as follows: |
|
1053 *} |
|
1054 |
|
1055 declare test[my_thms del] |
|
1056 |
|
1057 text {* After this, the theorem list is again: |
|
1058 |
|
1059 @{ML_response_fake [display,gray] |
|
1060 "!my_thms" |
|
1061 "[\"True\"]"} |
|
1062 |
|
1063 We used in this example two functions declared as @{ML Thm.declaration_attribute}, |
|
1064 but there can be any number of them. We just have to change the parser for reading |
|
1065 the arguments accordingly. |
|
1066 |
|
1067 However, as said at the beginning of this example, using references for storing theorems is |
|
1068 \emph{not} the received way of doing such things. The received way is to |
|
1069 start a ``data slot'' in a context by using the functor @{text GenericDataFun}: |
|
1070 *} |
|
1071 |
|
1072 ML {*structure Data = GenericDataFun |
|
1073 (type T = thm list |
|
1074 val empty = [] |
|
1075 val extend = I |
|
1076 fun merge _ = Thm.merge_thms) *} |
|
1077 |
|
1078 text {* |
|
1079 To use this data slot, you only have to change @{ML my_thms_add} and |
|
1080 @{ML my_thms_del} to: |
|
1081 *} |
|
1082 |
|
1083 ML{*val thm_add = Data.map o Thm.add_thm |
|
1084 val thm_del = Data.map o Thm.del_thm*} |
|
1085 |
|
1086 text {* |
|
1087 where @{ML Data.map} updates the data appropriately in the context. Since storing |
|
1088 theorems in a special list is such a common task, there is the functor @{text NamedThmsFun}, |
|
1089 which does most of the work for you. To obtain such a named theorem lists, you just |
|
1090 declare |
|
1091 *} |
|
1092 |
|
1093 ML{*structure FooRules = NamedThmsFun |
|
1094 (val name = "foo" |
|
1095 val description = "Rules for foo"); *} |
|
1096 |
|
1097 text {* |
|
1098 and set up the @{ML_struct FooRules} with the command |
|
1099 *} |
|
1100 |
|
1101 setup {* FooRules.setup *} |
|
1102 |
|
1103 text {* |
|
1104 This code declares a data slot where the theorems are stored, |
|
1105 an attribute @{text foo} (with the @{text add} and @{text del} options |
|
1106 for adding and deleting theorems) and an internal ML interface to retrieve and |
|
1107 modify the theorems. |
|
1108 |
|
1109 Furthermore, the facts are made available on the user-level under the dynamic |
|
1110 fact name @{text foo}. For example you can declare three lemmas to be of the kind |
|
1111 @{text foo} by: |
|
1112 *} |
|
1113 |
|
1114 lemma rule1[foo]: "A" sorry |
|
1115 lemma rule2[foo]: "B" sorry |
|
1116 lemma rule3[foo]: "C" sorry |
|
1117 |
|
1118 text {* and undeclare the first one by: *} |
|
1119 |
|
1120 declare rule1[foo del] |
|
1121 |
|
1122 text {* and query the remaining ones with: |
|
1123 |
|
1124 \begin{isabelle} |
|
1125 \isacommand{thm}~@{text "foo"}\\ |
|
1126 @{text "> ?C"}\\ |
|
1127 @{text "> ?B"} |
|
1128 \end{isabelle} |
|
1129 |
|
1130 On the ML-level the rules marked with @{text "foo"} can be retrieved |
|
1131 using the function @{ML FooRules.get}: |
|
1132 |
|
1133 @{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"} |
|
1134 |
|
1135 \begin{readmore} |
|
1136 For more information see @{ML_file "Pure/Tools/named_thms.ML"} and also |
|
1137 the recipe in Section~\ref{recipe:storingdata} about storing arbitrary |
|
1138 data. |
|
1139 \end{readmore} |
|
1140 |
|
1141 (FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?) |
963 |
1142 |
964 |
1143 |
965 \begin{readmore} |
1144 \begin{readmore} |
966 FIXME: @{ML_file "Pure/more_thm.ML"} @{ML_file "Pure/Isar/attrib.ML"} |
1145 FIXME: @{ML_file "Pure/more_thm.ML"} @{ML_file "Pure/Isar/attrib.ML"} |
967 \end{readmore} |
1146 \end{readmore} |
968 |
1147 *} |
969 |
1148 |
970 *} |
1149 |
971 |
1150 section {* Theories, Contexts and Local Theories (TBD) *} |
972 |
|
973 section {* Theories and Local Theories *} |
|
974 |
1151 |
975 text {* |
1152 text {* |
976 (FIXME: expand) |
1153 (FIXME: expand) |
977 |
1154 |
978 (FIXME: explain \isacommand{setup}) |
1155 (FIXME: explain \isacommand{setup}) |
990 |
1167 |
991 *} |
1168 *} |
992 |
1169 |
993 |
1170 |
994 |
1171 |
995 section {* Storing Theorems\label{sec:storing} *} |
1172 section {* Storing Theorems\label{sec:storing} (TBD) *} |
996 |
1173 |
997 text {* @{ML PureThy.add_thms_dynamic} *} |
1174 text {* @{ML PureThy.add_thms_dynamic} *} |
998 |
1175 |
999 |
1176 |
1000 |
1177 |
1001 (* FIXME: some code below *) |
1178 (* FIXME: some code below *) |
1002 |
1179 |
1003 (*<*) |
1180 (*<*) |
|
1181 (* |
1004 setup {* |
1182 setup {* |
1005 Sign.add_consts_i [("bar", @{typ "nat"},NoSyn)] |
1183 Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)] |
1006 *} |
1184 *} |
1007 |
1185 *) |
1008 lemma "bar = (1::nat)" |
1186 lemma "bar = (1::nat)" |
1009 oops |
1187 oops |
1010 |
1188 |
|
1189 (* |
1011 setup {* |
1190 setup {* |
1012 Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)] |
1191 Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)] |
1013 #> PureThy.add_defs false [((Binding.name "foo_def", |
1192 #> PureThy.add_defs false [((Binding.name "foo_def", |
1014 Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] |
1193 Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] |
1015 #> snd |
1194 #> snd |
1016 *} |
1195 *} |
1017 |
1196 *) |
|
1197 (* |
1018 lemma "foo = (1::nat)" |
1198 lemma "foo = (1::nat)" |
1019 apply(simp add: foo_def) |
1199 apply(simp add: foo_def) |
1020 done |
1200 done |
1021 |
1201 |
1022 thm foo_def |
1202 thm foo_def |
|
1203 *) |
1023 (*>*) |
1204 (*>*) |
1024 |
1205 |
1025 section {* Misc *} |
1206 section {* Pretty-Printing (TBD) *} |
|
1207 |
|
1208 text {* |
|
1209 @{ML Pretty.big_list}, |
|
1210 @{ML Pretty.brk}, |
|
1211 @{ML Pretty.block}, |
|
1212 @{ML Pretty.chunks} |
|
1213 *} |
|
1214 |
|
1215 section {* Misc (TBD) *} |
1026 |
1216 |
1027 ML {*DatatypePackage.get_datatype @{theory} "List.list"*} |
1217 ML {*DatatypePackage.get_datatype @{theory} "List.list"*} |
1028 |
1218 |
1029 end |
1219 end |