src/HOL/Decision_Procs/langford_data.ML
author hoelzl
Mon, 14 Mar 2011 14:37:41 +0100
changeset 41975 d47eabd80e59
parent 41226 adcb9a1198e7
child 51717 9e7d1c139569
permissions -rw-r--r--
simplified definition of open_extreal
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     1
signature LANGFORD_DATA =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     2
sig
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     3
  type entry
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     4
  val get: Proof.context -> simpset * (thm * entry) list
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     5
  val add: entry -> attribute 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     6
  val del: attribute
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     7
  val setup: theory -> theory
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     8
  val match: Proof.context -> cterm -> entry option
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
     9
end;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    10
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    11
structure Langford_Data: LANGFORD_DATA = 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    12
struct
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    13
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    14
(* data *)
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    15
type entry = {qe_bnds: thm, qe_nolb : thm , qe_noub: thm, 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    16
              gs : thm list, gst : thm, atoms: cterm list};
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    17
val eq_key = Thm.eq_thm;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    18
fun eq_data arg = eq_fst eq_key arg;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    19
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30722
diff changeset
    20
structure Data = Generic_Data
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30722
diff changeset
    21
(
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30722
diff changeset
    22
  type T = simpset * (thm * entry) list;
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    23
  val empty = (HOL_basic_ss, []);
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    24
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30722
diff changeset
    25
  fun merge ((ss1, es1), (ss2, es2)) : T =
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30722
diff changeset
    26
    (merge_ss (ss1, ss2), AList.merge eq_key (K true) (es1, es2));
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    27
);
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    28
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    29
val get = Data.get o Context.Proof;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    30
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    31
fun del_data key = apsnd (remove eq_data (key, []));
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    32
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    33
val del = Thm.declaration_attribute (Data.map o del_data);
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    34
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    35
fun add entry = 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    36
    Thm.declaration_attribute (fn key => fn context => context |> Data.map 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    37
      (del_data key #> apsnd (cons (key, entry))));
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    38
41226
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 37120
diff changeset
    39
val add_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.add_simp);
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    40
41226
adcb9a1198e7 clarified exports of structure Simplifier;
wenzelm
parents: 37120
diff changeset
    41
val del_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.del_simp);
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    42
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    43
fun match ctxt tm =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    44
  let
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    45
    fun match_inst
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    46
        {qe_bnds, qe_nolb, qe_noub, gs, gst, atoms} pat =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    47
       let
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    48
        fun h instT =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    49
          let
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    50
            val substT = Thm.instantiate (instT, []);
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    51
            val substT_cterm = Drule.cterm_rule substT;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    52
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    53
            val qe_bnds' = substT qe_bnds
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    54
            val qe_nolb' = substT qe_nolb
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    55
            val qe_noub' = substT qe_noub
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    56
            val gs' = map substT gs
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    57
            val gst' = substT gst
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    58
            val atoms' = map substT_cterm atoms
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    59
            val result = {qe_bnds = qe_bnds', qe_nolb = qe_nolb', 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    60
                          qe_noub = qe_noub', gs = gs', gst = gst', 
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    61
                          atoms = atoms'}
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    62
          in SOME result end
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    63
      in (case try Thm.match (pat, tm) of
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    64
           NONE => NONE
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    65
         | SOME (instT, _) => h instT)
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    66
      end;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    67
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    68
    fun match_struct (_,
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    69
        entry as ({atoms = atoms, ...}): entry) =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    70
      get_first (match_inst entry) atoms;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    71
  in get_first match_struct (snd (get ctxt)) end;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    72
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    73
(* concrete syntax *)
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    74
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    75
local
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    76
val qeN = "qe";
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    77
val gatherN = "gather";
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    78
val atomsN = "atoms"
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    79
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    80
val any_keyword =
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    81
  keyword qeN || keyword gatherN || keyword atomsN;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    82
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    83
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    84
val terms = thms >> map Drule.dest_term;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    85
in
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    86
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    87
val langford_setup =
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    88
  Attrib.setup @{binding langford}
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    89
    ((keyword qeN |-- thms)
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
    90
     -- (keyword gatherN |-- thms)
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    91
     -- (keyword atomsN |-- terms) >>
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    92
       (fn ((qes,gs), atoms) => 
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    93
       if length qes = 3 andalso length gs > 1 then
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    94
         let 
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    95
           val [q1,q2,q3] = qes
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    96
           val gst::gss = gs
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    97
           val entry = {qe_bnds = q1, qe_nolb = q2,
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    98
                        qe_noub = q3, gs = gss, gst = gst, atoms = atoms}
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
    99
         in add entry end
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   100
       else error "Need 3 theorems for qe and at least one for gs"))
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   101
    "Langford data";
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   102
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   103
end;
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   104
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   105
(* theory setup *)
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   106
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   107
val setup =
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   108
  langford_setup #>
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 25979
diff changeset
   109
  Attrib.setup @{binding langfordsimp} (Attrib.add_del add_simp del_simp) "Langford simpset";
23907
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   110
cca404a62173 Context data for QE in DLO (Langford's algorithm)
chaieb
parents:
diff changeset
   111
end;