677 @{ML_response [display,gray] |
677 @{ML_response [display,gray] |
678 "let |
678 "let |
679 val input = filtered_input |
679 val input = filtered_input |
680 \"foo::\\\"int \<Rightarrow> bool\\\" and bar::nat (\\\"BAR\\\" 100) and blonk\" |
680 \"foo::\\\"int \<Rightarrow> bool\\\" and bar::nat (\\\"BAR\\\" 100) and blonk\" |
681 in |
681 in |
682 parse OuterParse.fixes input |
682 parse OuterParse.fixes input |
683 end" |
683 end" |
684 "([(foo, SOME \"\\^E\\^Ftoken\\^Eint \<Rightarrow> bool\\^E\\^F\\^E\", NoSyn), |
684 "([(foo, SOME \"\\^E\\^Ftoken\\^Eint \<Rightarrow> bool\\^E\\^F\\^E\", NoSyn), |
685 (bar, SOME \"\\^E\\^Ftoken\\^Enat\\^E\\^F\\^E\", Mixfix (\"BAR\", [], 100)), |
685 (bar, SOME \"\\^E\\^Ftoken\\^Enat\\^E\\^F\\^E\", Mixfix (\"BAR\", [], 100)), |
686 (blonk, NONE, NoSyn)],[])"} |
686 (blonk, NONE, NoSyn)],[])"} |
687 *} |
687 *} |
697 \begin{readmore} |
697 \begin{readmore} |
698 The datastructre for sytax annotations is defined in @{ML_file "Pure/Syntax/mixfix.ML"}. |
698 The datastructre for sytax annotations is defined in @{ML_file "Pure/Syntax/mixfix.ML"}. |
699 \end{readmore} |
699 \end{readmore} |
700 |
700 |
701 Lines 3 to 7 in the function @{ML spec_parser} implement the parser for a |
701 Lines 3 to 7 in the function @{ML spec_parser} implement the parser for a |
702 list of introduction rules, that is propositions with theorem |
702 list of introduction rules, that is propositions with theorem annotations |
703 annotations. The introduction rules are propositions parsed by @{ML |
703 such as rule names and attributes. The introduction rules are propositions |
704 OuterParse.prop}. However, they can include an optional theorem name plus |
704 parsed by @{ML OuterParse.prop}. However, they can include an optional |
705 some attributes. For example |
705 theorem name plus some attributes. For example |
706 |
706 |
707 @{ML_response [display,gray] "let |
707 @{ML_response [display,gray] "let |
708 val input = filtered_input \"foo_lemma[intro,dest!]:\" |
708 val input = filtered_input \"foo_lemma[intro,dest!]:\" |
709 val ((name, attrib), _) = parse (SpecParse.thm_name \":\") input |
709 val ((name, attrib), _) = parse (SpecParse.thm_name \":\") input |
710 in |
710 in |
764 at all. We shall name this command \isacommand{foobar}. On the ML-level it can be |
762 at all. We shall name this command \isacommand{foobar}. On the ML-level it can be |
765 defined as: |
763 defined as: |
766 *} |
764 *} |
767 |
765 |
768 ML{*let |
766 ML{*let |
769 val do_nothing = Scan.succeed (Toplevel.theory I) |
767 val do_nothing = Scan.succeed (LocalTheory.theory I) |
770 val kind = OuterKeyword.thy_decl |
768 val kind = OuterKeyword.thy_decl |
771 in |
769 in |
772 OuterSyntax.command "foobar" "description of foobar" kind do_nothing |
770 OuterSyntax.local_theory "foobar" "description of foobar" kind do_nothing |
773 end *} |
771 end *} |
774 |
772 |
775 text {* |
773 text {* |
776 The crucial function @{ML OuterSyntax.command} expects a name for the command, a |
774 The crucial function @{ML OuterSyntax.local_theory} expects a name for the command, a |
777 short description, a kind indicator (which we will explain later on more thoroughly) and a |
775 short description, a kind indicator (which we will explain later more thoroughly) and a |
778 parser producing a top-level transition function (its purpose will also explained |
776 parser producing a local theory transition (its purpose will also explained |
779 later). |
777 later). |
780 |
778 |
781 While this is everything you have to do on the ML-level, you need a keyword |
779 While this is everything you have to do on the ML-level, you need a keyword |
782 file that can be loaded by ProofGeneral. This is to enable ProofGeneral to |
780 file that can be loaded by ProofGeneral. This is to enable ProofGeneral to |
783 recognise \isacommand{foobar} as a command. Such a keyword file can be |
781 recognise \isacommand{foobar} as a command. Such a keyword file can be |
801 \isacommand{imports}~@{text Main}\\ |
799 \isacommand{imports}~@{text Main}\\ |
802 \isacommand{begin}\\ |
800 \isacommand{begin}\\ |
803 \isacommand{ML}~@{text "\<verbopen>"}\\ |
801 \isacommand{ML}~@{text "\<verbopen>"}\\ |
804 @{ML |
802 @{ML |
805 "let |
803 "let |
806 val do_nothing = Scan.succeed (Toplevel.theory I) |
804 val do_nothing = Scan.succeed (LocalTheory.theory I) |
807 val kind = OuterKeyword.thy_decl |
805 val kind = OuterKeyword.thy_decl |
808 in |
806 in |
809 OuterSyntax.command \"foobar\" \"description of foobar\" kind do_nothing |
807 OuterSyntax.local_theory \"foobar\" \"description of foobar\" kind do_nothing |
810 end"}\\ |
808 end"}\\ |
811 @{text "\<verbclose>"}\\ |
809 @{text "\<verbclose>"}\\ |
812 \isacommand{end} |
810 \isacommand{end} |
813 \end{graybox} |
811 \end{graybox} |
814 \caption{\small The file @{text "Command.thy"} is necessary for generating a log |
812 \caption{The file @{text "Command.thy"} is necessary for generating a log |
815 file. This log file enables Isabelle to generate a keyword file containing |
813 file. This log file enables Isabelle to generate a keyword file containing |
816 the command \isacommand{foobar}.\label{fig:commandtheory}} |
814 the command \isacommand{foobar}.\label{fig:commandtheory}} |
817 \end{figure} |
815 \end{figure} |
818 %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% |
816 %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% |
819 |
817 |
907 inside the tracing buffer. |
905 inside the tracing buffer. |
908 |
906 |
909 The crucial part of a command is the function that determines the behaviour |
907 The crucial part of a command is the function that determines the behaviour |
910 of the command. In the code above we used a ``do-nothing''-function, which |
908 of the command. In the code above we used a ``do-nothing''-function, which |
911 because of @{ML Scan.succeed} does not parse any argument, but immediately |
909 because of @{ML Scan.succeed} does not parse any argument, but immediately |
912 returns the simple toplevel function @{ML "Toplevel.theory I"}. We can |
910 returns the simple function @{ML "LocalTheory.theory I"}. We can |
913 replace this code by a function that first parses a proposition (using the |
911 replace this code by a function that first parses a proposition (using the |
914 parser @{ML OuterParse.prop}), then prints out the tracing |
912 parser @{ML OuterParse.prop}), then prints out the tracing |
915 information (using a new top-level function @{text trace_top_lvl}) and |
913 information (using a new function @{text trace_prop}) and |
916 finally does nothing. For this you can write: |
914 finally does nothing. For this you can write: |
917 *} |
915 *} |
918 |
916 |
919 ML{*let |
917 ML{*let |
920 fun trace_top_lvl str = |
918 fun trace_prop str = |
921 LocalTheory.theory (fn lthy => (tracing str; lthy)) |
919 LocalTheory.theory (fn lthy => (tracing str; lthy)) |
922 |
920 |
923 val trace_prop = OuterParse.prop >> trace_top_lvl |
921 val trace_prop_parser = OuterParse.prop >> trace_prop |
924 |
|
925 val kind = OuterKeyword.thy_decl |
922 val kind = OuterKeyword.thy_decl |
926 in |
923 in |
927 OuterSyntax.local_theory "foobar" "traces a proposition" kind trace_prop |
924 OuterSyntax.local_theory "foobar" "traces a proposition" |
|
925 kind trace_prop_parser |
928 end *} |
926 end *} |
929 |
|
930 |
927 |
931 text {* |
928 text {* |
932 Now you can type |
929 Now you can type |
933 |
930 |
934 \begin{isabelle} |
931 \begin{isabelle} |
936 @{text "> \"True \<and> False\""} |
933 @{text "> \"True \<and> False\""} |
937 \end{isabelle} |
934 \end{isabelle} |
938 |
935 |
939 and see the proposition in the tracing buffer. |
936 and see the proposition in the tracing buffer. |
940 |
937 |
941 Note that so far we used @{ML thy_decl in OuterKeyword} as the kind indicator |
938 Note that so far we used @{ML thy_decl in OuterKeyword} as the kind |
942 for the command. This means that the command finishes as soon as the |
939 indicator for the command. This means that the command finishes as soon as |
943 arguments are processed. Examples of this kind of commands are |
940 the arguments are processed. Examples of this kind of commands are |
944 \isacommand{definition} and \isacommand{declare}. In other cases, |
941 \isacommand{definition} and \isacommand{declare}. In other cases, commands |
945 commands are expected to parse some arguments, for example a proposition, |
942 are expected to parse some arguments, for example a proposition, and then |
946 and then ``open up'' a proof in order to prove the proposition (for example |
943 ``open up'' a proof in order to prove the proposition (for example |
947 \isacommand{lemma}) or prove some other properties (for example |
944 \isacommand{lemma}) or prove some other properties (for example |
948 \isacommand{function}). To achieve this kind of behaviour, you have to use the kind |
945 \isacommand{function}). To achieve this kind of behaviour, you have to use |
949 indicator @{ML thy_goal in OuterKeyword}. Note, however, once you change the |
946 the kind indicator @{ML thy_goal in OuterKeyword} and the function @{ML |
950 ``kind'' of a command from @{ML thy_decl in OuterKeyword} to @{ML thy_goal in OuterKeyword} |
947 "local_theory_to_proof" in OuterSyntax} to set up the command. Note, |
951 then the keyword file needs to be re-created! |
948 however, once you change the ``kind'' of a command from @{ML thy_decl in |
|
949 OuterKeyword} to @{ML thy_goal in OuterKeyword} then the keyword file needs |
|
950 to be re-created! |
952 |
951 |
953 Below we change \isacommand{foobar} so that it takes a proposition as |
952 Below we change \isacommand{foobar} so that it takes a proposition as |
954 argument and then starts a proof in order to prove it. Therefore in Line 13, |
953 argument and then starts a proof in order to prove it. Therefore in Line 13, |
955 we set the kind indicator to @{ML thy_goal in OuterKeyword}. |
954 we set the kind indicator to @{ML thy_goal in OuterKeyword}. |
956 *} |
955 *} |
957 |
956 |
958 ML%linenosgray{*let |
957 ML%linenosgray{*let |
959 fun set_up_thm str ctxt = |
958 fun prove_prop str ctxt = |
960 let |
959 let |
961 val prop = Syntax.read_prop ctxt str |
960 val prop = Syntax.read_prop ctxt str |
962 in |
961 in |
963 Proof.theorem_i NONE (K I) [[(prop,[])]] ctxt |
962 Proof.theorem_i NONE (K I) [[(prop,[])]] ctxt |
964 end; |
963 end; |
965 |
964 |
966 val prove_prop = OuterParse.prop >> |
965 val prove_prop_parser = OuterParse.prop >> prove_prop |
967 (fn str => Toplevel.print o |
|
968 Toplevel.local_theory_to_proof NONE (set_up_thm str)) |
|
969 |
|
970 val kind = OuterKeyword.thy_goal |
966 val kind = OuterKeyword.thy_goal |
971 in |
967 in |
972 OuterSyntax.command "foobar" "proving a proposition" kind prove_prop |
968 OuterSyntax.local_theory_to_proof "foobar" "proving a proposition" |
|
969 kind prove_prop_parser |
973 end *} |
970 end *} |
974 |
971 |
975 text {* |
972 text {* |
976 The function @{text set_up_thm} in Lines 2 to 7 takes a string (the proposition to be |
973 The function @{text prove_prop} in Lines 2 to 7 takes a string (the proposition to be |
977 proved) and a context as argument. The context is necessary in order to be able to use |
974 proved) and a context as argument. The context is necessary in order to be able to use |
978 @{ML Syntax.read_prop}, which converts a string into a proper proposition. |
975 @{ML Syntax.read_prop}, which converts a string into a proper proposition. |
979 In Line 6 the function @{ML Proof.theorem_i} starts the proof for the |
976 In Line 6 the function @{ML Proof.theorem_i} starts the proof for the |
980 proposition. Its argument @{ML NONE} stands for a locale (which we chose to |
977 proposition. Its argument @{ML NONE} stands for a locale (which we chose to |
981 omit); the argument @{ML "(K I)"} stands for a function that determines what |
978 omit); the argument @{ML "(K I)"} stands for a function that determines what |
982 should be done with the theorem once it is proved (we chose to just forget |
979 should be done with the theorem once it is proved (we chose to just forget |
983 about it). Lines 9 to 11 contain the parser for the proposition. |
980 about it). Line 9 contains the parser for the proposition. |
984 |
981 |
985 If you now type \isacommand{foobar}~@{text [quotes] "True \<and> True"}, you obtain the following |
982 If you now type \isacommand{foobar}~@{text [quotes] "True \<and> True"}, you obtain the following |
986 proof state: |
983 proof state |
987 |
984 |
988 \begin{isabelle} |
985 \begin{isabelle} |
989 \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\ |
986 \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\ |
990 @{text "goal (1 subgoal):"}\\ |
987 @{text "goal (1 subgoal):"}\\ |
991 @{text "1. True \<and> True"} |
988 @{text "1. True \<and> True"} |
992 \end{isabelle} |
989 \end{isabelle} |
993 |
990 |
994 and you can build the proof |
991 and you can build the following proof |
995 |
992 |
996 \begin{isabelle} |
993 \begin{isabelle} |
997 \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\ |
994 \isacommand{foobar}~@{text [quotes] "True \<and> True"}\\ |
998 \isacommand{apply}@{text "(rule conjI)"}\\ |
995 \isacommand{apply}@{text "(rule conjI)"}\\ |
999 \isacommand{apply}@{text "(rule TrueI)+"}\\ |
996 \isacommand{apply}@{text "(rule TrueI)+"}\\ |
1000 \isacommand{done} |
997 \isacommand{done} |
1001 \end{isabelle} |
998 \end{isabelle} |
1002 |
|
1003 |
|
1004 |
|
1005 (FIXME What do @{ML "Toplevel.theory"} |
|
1006 @{ML "Toplevel.print"} |
|
1007 @{ML Toplevel.local_theory} do?) |
|
1008 |
999 |
1009 (FIXME read a name and show how to store theorems) |
1000 (FIXME read a name and show how to store theorems) |
1010 *} |
1001 *} |
1011 |
1002 |
1012 section {* Methods (TBD) *} |
1003 section {* Methods (TBD) *} |