# HG changeset patch # User wenzelm # Date 1417722998 -3600 # Node ID a0a05a4edb36c950e0df27eb5da5878ed2158d03 # Parent da2fef2faa83cd2bdeac4f6790cefbcebb384fd4 more examples; diff -r da2fef2faa83 -r a0a05a4edb36 src/Doc/Implementation/Integration.thy --- a/src/Doc/Implementation/Integration.thy Thu Dec 04 20:45:11 2014 +0100 +++ b/src/Doc/Implementation/Integration.thy Thu Dec 04 20:56:38 2014 +0100 @@ -150,6 +150,10 @@ \end{description} \ +text %mlex \The file @{"file" "~~/src/HOL/ex/Commands.thy"} shows some example +Isar command definitions, with the all-important theory header declarations +for outer syntax keywords.\ + section \Theory loader database\ diff -r da2fef2faa83 -r a0a05a4edb36 src/HOL/ROOT --- a/src/HOL/ROOT Thu Dec 04 20:45:11 2014 +0100 +++ b/src/HOL/ROOT Thu Dec 04 20:56:38 2014 +0100 @@ -539,6 +539,7 @@ "~~/src/HOL/Library/Transitive_Closure_Table" Cartouche_Examples theories + Commands Adhoc_Overloading_Examples Iff_Oracle Coercion_Examples diff -r da2fef2faa83 -r a0a05a4edb36 src/HOL/ex/Commands.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Commands.thy Thu Dec 04 20:56:38 2014 +0100 @@ -0,0 +1,76 @@ +(* Title: HOL/ex/Commands.thy + Author: Makarius +*) + +section \Some Isar command definitions\ + +theory Commands +imports Main +keywords + "print_test" :: diag and + "global_test" :: thy_decl and + "local_test" :: thy_decl +begin + +subsection \Diagnostic command: no state change\ + +ML \ + Outer_Syntax.command @{command_spec "print_test"} "print term test" + (Parse.term >> (fn s => Toplevel.keep (fn st => + let + val ctxt = Toplevel.context_of st; + val t = Syntax.read_term ctxt s; + in Pretty.writeln (Syntax.pretty_term ctxt t) end))); +\ + +print_test x +print_test "\x. x = a" + + +subsection \Old-style global theory declaration\ + +ML \ + Outer_Syntax.command @{command_spec "global_test"} "test constant declaration" + (Parse.binding >> (fn b => Toplevel.theory (fn thy => + let + val thy' = Sign.add_consts [(b, @{typ 'a}, NoSyn)] thy; + in thy' end))); +\ + +global_test a +global_test b +print_test a + + +subsection \Local theory specification\ + +ML \ + Outer_Syntax.local_theory @{command_spec "local_test"} "test local definition" + (Parse.binding -- (@{keyword "="} |-- Parse.term) >> (fn (b, s) => fn lthy => + let + val t = Syntax.read_term lthy s; + val (def, lthy') = Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), t)) lthy; + in lthy' end)); +\ + +local_test true = True +print_test true +thm true_def + +local_test identity = "\x. x" +print_test "identity x" +thm identity_def + +context fixes x y :: nat +begin + +local_test test = "x + y" +print_test test +thm test_def + +end + +print_test "test 0 1" +thm test_def + +end