Commit Graph

51 Commits

Author SHA1 Message Date
Nicolas Méric d2a6106be5 Fix the record generation in class implementation
- Fix the generation of the record associated with
  a class and used for the logic.
  The old implementation generated a new attribute
  for each attribute defined by a subclass,
  even the ones that were overriding ones of the superclass.
  The new implementation generates the attributes of the subclass
  which are not overriding ones.
  Warning:
  It implies that overridden attributes in a subclass are not
  new attributes added to the theory context.
  So the base name of an attribute will refer to the attribute
  of the last declared class where it is defined.
  If ones wants to refer to atttributes, one should use
  long names, even in the invariants of a subclass definition
  which overrides the attribute used in the invariant.
  For example,
  in ~~/src/ontologies/scholarly_paper/scholarly_paper.thy:

  doc_class technical = text_section +
     definition_list :: "string list" <=  "[]"
     status          :: status <= "description"
     formal_results  :: "thm list"
     invariant L1    :: "λσ::technical. the (level σ) > 0"

  type_synonym tc = technical (* technical content *)

  doc_class example  = text_section +
     referentiable   :: bool <= True
     status          :: status <= "description"
     short_name      :: string <= "''''"

  doc_class math_content = tc +
     referentiable :: bool <= True
     short_name    :: string <= "''''"
     status        :: status <= "semiformal"
     mcc           :: "math_content_class" <= "thm"
     invariant s1  :: "λ σ::math_content. ¬referentiable σ ⟶ short_name σ = ''''"
     invariant s2  :: "λ σ::math_content. technical.status σ = semiformal"

  The class math_content overrride the attribute status
  of the class technical, by using the type synonym tc,
  but the base name of this attribute refers
  to the attribute of the class example where it is last defined
  and not just overridden.
  So in the invariant s2 of the class math_content,
  we must use the long name of the attribute,
  i.e. the base name "status" with its qualifier which refers
  to the superclass where it is defined, the class technical.

  Type synonyms as qualifiers are not yet supported.
- Qualify classes that only override attributes of their superclass
  as vitual classes by adding a virtual attribute.
  This attribute is used to discriminate virtual classes and generate
  an adequate make function to initialize their associated record.
  The implementation uses an hidden attribute (the tag_attribute)
  to force the virtual class to be concrete or the logic
  by having a full new record definition associated with it.
  For example:

  doc_class W =
    a::"int" <= "1"

  doc_class X = W +
    a::"int" <= "2"

  The class X is tagged as a virtual class and
  the record make functions of the classes W and X are:

  W.make W_tag_attribute W_a
  X.make X_tag_attribute X_a X_tag_attribute

  So a record definition is added to the theory context for each class,
  even though a virtual class only overrides
  attributes of its superclass.
  This behavior allows us to support definitions of new default values
  for attributes in the subclass, as shown in the example.
- Factorize make name components
- Use Record name components instead of strings to refer to Record
  components
- Fix typos
2021-12-07 10:04:47 +01:00
Nicolas Méric 08c101c544 Implement built-ins referential equivalence
- Add a first implementation of a referential equivalence
  for the built-ins term annotations (TA)
- Some built-ins remain as unspecified constants:
  - the docitem TA offers a way to check the reference of
    class instances without checking the instances type.
    It must be avoided for certification
  - the termrepr TA is left as an unspecified constant for now.
    A major refactoring of code should be done to enable
    referential equivalence for termrepr, by changing the dependency
    between the Isa_DOF theory and the Assert theory.
    The assert_cmd function in Assert should use the value* command
    functions, which make the elaboration of the term
    referenced by the TA before passing it to the evaluator
- Update the Evaluation test theory to test the referential equivalence
  and expose some of  current implementation limitations
- Add a warning about the docitem TA in the TermAntiquotations theory
2021-11-09 08:55:02 +01:00
Nicolas Méric 6ac1445147 Change the implementation of the tag attribute
The philosophy is for the tag attribute to be unique
for each class.
So this commit updates the implementation of this attribute
to match the philosophy.
The previous implementation associated a tag attribute
with a class but also with each super-class of this class
up to the top (default) class "text".
Now a class with super-classes has only one tag attribute.
2021-11-08 10:44:29 +01:00
Nicolas Méric 664aede4c0 First draft of the value* command implementation
Add a command value*
- The value* command uses the same code as the value command
  and adds the possibility to evaluate
  Term Annotation Antiquotations (TA)
  with the help of the DOF_core.transduce_term_global function.
  The DOF_core.transduce_term_global function,
  in addition to the validation of a term
  (also called a second level type checking),
  is now able to make a so called elaboration:
  it will construct the term referenced by a TA before
  passing it to the evaluator.
- For a term to be evaluated, it must not be contain
  the "undefined" constant whose evaluation always fails.
  (See the code generation documentation).
  Furthermore, the instance class generation is updated in such a way
  that each of its attributes is initialized with a free variable
  whose name shows to the final user that this attribute
  is not initialized.
  It implies that an instance class evaluation will be pass
  to the normalization by evaluation (nbe) evaluator by default
  if the final user does not specify a class instance entirely,
  i.e. by specifying each attribute of the instance.
  This choice is considered a decent compromise, considering
  the speed and trustworthiness of the nbe evaluator.
  (See the article
  A Compiled Implementation of Normalization by Evaluation from 2008)
- Update the ISA transformer tab to add a function
  which is used for the elaboration of the term referenced by the TA.to pass
- Add a first really basic draft of the implementation
  of the elaboration of the built-ins TA and of an instance class:
  - For the built-ins, the term referenced by the TA is returned
    as it is;
  - For an instance class, the value of the instance is returned.
- Make the tag attribute global by moving it to DOF_core structure
- Add a first draft for some evaluation tests
  and expose the limitations of the current implementation
  in Evaluation.thy
2021-11-08 10:38:11 +01:00
Burkhart Wolff 4420084d52 restructuring command-syntax doc_class 2021-09-29 14:21:13 +02:00
Nicolas Méric 2c01a7118b Add term* cmd and term antiquotations for classes
- Add a term antiquotation for document classes
  and add the term* command which mimics the classical term command
  and adds the possibility to use a term antiquotation
  which references document classes:
  one can use @{A ''A_instance''} to reference
  an instance of the class A in a term* command.
- Reuse and update the ML_isa_check_generic visitor pattern
  to add the function which checks the class instance of a class,
  used by the term antiquotation for document classes.
  Also, the update_isa functions now expect long name
  (See builtin term antiquotations setup)
- The merge of ISA_transformer_tab has been update to avoid conflicts.
  Indeed, the merge is ultra-critical: the transformer tabs were
  just extended by letting *the first* entry
  with the same long-name win.
  Since the range is a (call-back) function, a comparison on its content
  is impossible and some choice has to be made.
  An alternative may be to use Symtab.join
- As classes names as constants are already bound to
  "doc_class Regular_Exp.rexp" constants by add_doc_class_cmd function,
  we use a prefix "doc_class_" when adding
  document classes term antiquotations
2021-06-01 17:32:45 +02:00
Burkhart Wolff 3b21df199b addded docitem ML antiquotation. (Kleine Fingeruebung). 2021-04-21 20:24:06 +02:00
Achim D. Brucker 06dddeacf5 Porting to Isabelle 2021. 2021-03-10 22:04:09 +00:00
Burkhart Wolff 396ef1d477 More content in 4, better tree printing. 2021-01-01 21:23:21 +01:00
Burkhart Wolff 242bb536bc Reorganization Chap 4. <4.3.2 2020-12-30 23:07:19 +01:00
Burkhart Wolff 04f0cc7f5c Reorganization: Pushed Macro Core Mechanism into the DOF Core; adapted the RefMan accordingly. 2020-12-30 12:47:54 +01:00
Burkhart Wolff 8771d8581b default class checking bug fixed; new attributes for default classes in ontological macros Definition* Theorem* Lemma* 2020-12-01 23:18:13 +01:00
Burkhart Wolff 2ecb62a80e added Lemma*, Theorem* and Definition* support. Bug: referencing does not work. 2020-11-04 15:55:43 +01:00
Burkhart Wolff da0f3e63f1 more steps to reform document macro mechanism 2020-11-04 13:13:24 +01:00
Burkhart Wolff bad7dfc2ef new set of macros : author* and abstract* --- not working yet 2020-11-03 19:00:33 +01:00
Burkhart Wolff e59ac46299 removed SI --- went to AFP 2020-11-02 14:32:41 +01:00
Burkhart Wolff 7a768cfdeb versatile 2020-08-26 08:43:39 +02:00
Burkhart Wolff 338bb7d4a4 Code cleanup. 2020-08-25 11:59:10 +02:00
Burkhart Wolff a792cc79d2 was lucky to solve a deep bug in standard antiquotation evaluation inside text* soon. 2020-08-25 11:11:38 +02:00
Burkhart Wolff 8002ec31bb cleanups after discussion 2020-08-24 14:36:22 +02:00
Burkhart Wolff 7cb6577797 solved presentation bug (brown) and eliminated some code dups 2020-08-24 11:33:32 +02:00
Burkhart Wolff 1470776428 slight correction of the template, and addition of SML template instance in DOF-technical_report. Does not work for test-case in 05_Implementation (Commented out) 2020-06-24 13:11:26 +02:00
Burkhart Wolff ef93285ec7 added a little useful template generation command 2020-06-23 14:02:04 +02:00
Burkhart Wolff 016a9e6454 Merge branch 'master' of https://git.logicalhacking.com/Isabelle_DOF/Isabelle_DOF 2020-06-22 17:42:48 +02:00
Burkhart Wolff 7e2224859e mmm 2020-06-22 17:42:40 +02:00
Burkhart Wolff 4717925eea Zwischenzustand OoO Generation 2020-06-16 09:08:36 +02:00
Burkhart Wolff 0f9b6731af replaced structure with legecy code: Pure_Syn_Ext. 2020-06-12 15:00:49 +02:00
Burkhart Wolff 9b0c2cdcd8 added support for defn, lem, thm short-calls. 2020-05-19 17:32:25 +02:00
Burkhart Wolff fa931b45e2 re-localization of onto macros. Tested. 2020-04-23 18:30:46 +02:00
Burkhart Wolff 8328626fa4 Restructuring library prep. 2020-04-23 16:08:05 +02:00
Burkhart Wolff 2e0d88a3f7 restructuring of COL, scholarly_paper, etc. Facturong out Macros. 2020-04-22 15:31:47 +02:00
Burkhart Wolff 88d4b7674e support along AMS style for mcc. 2020-04-14 14:44:42 +02:00
Burkhart Wolff 6cf5708d93 added macro_def mechanism. Bug: Type qualification necessary 2020-04-12 21:11:54 +02:00
Burkhart Wolff e642243847 added macrodef - expand mechanism 2020-04-10 18:30:33 +02:00
Burkhart Wolff 0c4a5a5fea eliminating deprecated syntax 2020-04-09 23:58:58 +02:00
Burkhart Wolff f1b376d4b6 added support for math_content-class in scholarly_paper in Knuth's Urschleim. 2020-04-09 17:25:09 +02:00
Burkhart Wolff aa4e1acf84 Added invariants - and changes of invariant syntax.
Modified scholarly_paper onto wrt to future concepts
of referential semi_formal items (according to discussion
with Achim).
2020-04-08 23:29:15 +02:00
Burkhart Wolff 9035c46023 syntax and 1st level type-checking of invariants 2020-02-21 19:23:51 +01:00
Burkhart Wolff cc98979f43 more on class_id synonyms 2020-02-21 16:33:28 +01:00
Burkhart Wolff 3faf3102ee First version with some places where type_synonyms were used to identify doc_classes 2020-02-21 15:39:50 +01:00
Burkhart Wolff 9df43f0085 various changes of the DOF-core interface: read_cid. Preparations for type_synonyms for cids. (unfinished). Updated scholarly_paper onto 2020-02-20 13:30:51 +01:00
Burkhart Wolff 5b0db2efb1 New Regrouping in the scholarly Onto + LaTeX support. Tested. 2020-02-19 18:13:33 +01:00
Burkhart Wolff c5b5f994ef added onto markup support for definitions and examples in scholarly_paper lncs style 2020-02-17 23:07:34 +01:00
Burkhart Wolff fe8a6c5c87 refinements of the technical class; added the document antiquotation doc_class; some experiments in SI. 2020-02-13 11:17:20 +01:00
Burkhart Wolff cb1ead378a added new sections in CommentedIsabelle concerning definitions and internal proofs 2019-12-17 13:25:43 +01:00
Achim D. Brucker 60ebbbe12c Updated license information. 2019-08-15 15:09:55 +01:00
Burkhart Wolff 6e44230efb Updated COL section ... 2019-08-15 11:30:42 +02:00
Burkhart Wolff b5fe2d9085 Solution to the assert - Bug : stronger checks in doc_class that reject correctly constructed, but lexically illegal long_names for doc_classes. 2019-08-14 17:22:55 +02:00
Achim D. Brucker 332daa1ebb Removed outdated and unsupported gen_sty_template. 2019-08-11 22:26:17 +01:00
Achim D. Brucker df12a32624 Added missing space to warning messages. 2019-08-04 08:24:01 +01:00