src/Doc/Codegen/Setup.thy
author wenzelm
Mon, 26 Nov 2012 10:37:05 +0100
changeset 50210 747db833fbf7
parent 48985 5386df44a037
child 51162 310b94ed1815
permissions -rw-r--r--
no special treatment of control_reset, in accordance to other control styles;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     1
theory Setup
39066
4517a4049588 more canonical theory setup
haftmann
parents: 38798
diff changeset
     2
imports
4517a4049588 more canonical theory setup
haftmann
parents: 38798
diff changeset
     3
  Complex_Main
46517
9d2e682a68eb tuned import
haftmann
parents: 43564
diff changeset
     4
  "~~/src/HOL/Library/Dlist"
9d2e682a68eb tuned import
haftmann
parents: 43564
diff changeset
     5
  "~~/src/HOL/Library/RBT"
9d2e682a68eb tuned import
haftmann
parents: 43564
diff changeset
     6
  "~~/src/HOL/Library/Mapping"
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     7
begin
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
     8
48951
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
     9
(* FIXME avoid writing into source directory *)
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    10
ML {*
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    11
  Isabelle_System.mkdirs (Path.append (Thy_Load.master_directory @{theory}) (Path.basic "examples"))
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    12
*}
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    13
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    14
ML_file "../antiquote_setup.ML"
b9238cbcdd41 more standard document preparation within session context;
wenzelm
parents: 48891
diff changeset
    15
ML_file "../more_antiquote.ML"
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 46517
diff changeset
    16
38503
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    17
setup {*
43564
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43271
diff changeset
    18
  Antiquote_Setup.setup #>
9864182c6bad document antiquotations are managed as theory data, with proper name space and entity markup;
wenzelm
parents: 43271
diff changeset
    19
  More_Antiquote.setup #>
38503
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    20
let
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    21
  val typ = Simple_Syntax.read_typ;
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    22
in
42293
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    23
  Sign.del_modesyntax_i (Symbol.xsymbolsN, false)
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    24
   [("_constrain", typ "logic => type => logic", Mixfix ("_\<Colon>_", [4, 0], 3)),
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    25
    ("_constrain", typ "prop' => type => prop'", Mixfix ("_\<Colon>_", [4, 0], 3))] #>
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    26
  Sign.add_modesyntax_i (Symbol.xsymbolsN, false)
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    27
   [("_constrain", typ "logic => type => logic", Mixfix ("_ \<Colon>  _", [4, 0], 3)),
6cca0343ea48 renamed sprop "prop#" to "prop'" -- proper identifier;
wenzelm
parents: 42288
diff changeset
    28
    ("_constrain", typ "prop' => type => prop'", Mixfix ("_ \<Colon> _", [4, 0], 3))]
38503
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    29
end
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    30
*}
7115853eaf8a pretty constraint syntax; tuned theory imports
haftmann
parents: 38460
diff changeset
    31
34071
93bfbb557e2e default_code_width is now proper theory data
haftmann
parents: 32833
diff changeset
    32
setup {* Code_Target.set_default_code_width 74 *}
93bfbb557e2e default_code_width is now proper theory data
haftmann
parents: 32833
diff changeset
    33
42669
04dfffda5671 more conventional naming scheme: names_long, names_short, names_unique;
wenzelm
parents: 42358
diff changeset
    34
declare [[names_unique = false]]
28213
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
    35
b52f9205a02d New outline for codegen tutorial -- draft
haftmann
parents:
diff changeset
    36
end
46517
9d2e682a68eb tuned import
haftmann
parents: 43564
diff changeset
    37