src/HOL/Tools/inductive_realizer.ML
author boehmes
Mon, 08 Nov 2010 12:13:44 +0100
changeset 40424 7550b2cba1cb
parent 39557 fe5722fce758
child 42361 23f352990944
permissions -rw-r--r--
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Tools/inductive_realizer.ML
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     3
36043
d149c3886e7e removed dead code; fixed typo
krauss
parents: 35845
diff changeset
     4
Program extraction from proofs involving inductive predicates:
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 28965
diff changeset
     5
Realizers for induction and elimination rules.
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     6
*)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     7
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     8
signature INDUCTIVE_REALIZER =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
     9
sig
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    10
  val add_ind_realizers: string -> string list -> theory -> theory
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18358
diff changeset
    11
  val setup: theory -> theory
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    12
end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    13
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    14
structure InductiveRealizer : INDUCTIVE_REALIZER =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    15
struct
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    16
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33669
diff changeset
    17
(* FIXME: Local_Theory.note should return theorems with proper names! *)  (* FIXME ?? *)
22606
962f824c2df9 - Tried to make name_of_thm more robust against changes of the
berghofe
parents: 22596
diff changeset
    18
fun name_of_thm thm =
28800
48f7bfebd31d name_of_thm: Proofterm.fold_proof_atoms;
wenzelm
parents: 28328
diff changeset
    19
  (case Proofterm.fold_proof_atoms false (fn PThm (_, ((name, _, _), _)) => cons name | _ => I)
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28800
diff changeset
    20
      [Thm.proof_of thm] [] of
28800
48f7bfebd31d name_of_thm: Proofterm.fold_proof_atoms;
wenzelm
parents: 28328
diff changeset
    21
    [name] => name
32091
30e2ffbba718 proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents: 31986
diff changeset
    22
  | _ => error ("name_of_thm: bad proof of theorem\n" ^ Display.string_of_thm_without_context thm));
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    23
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    24
fun prf_of thm =
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26535
diff changeset
    25
  let
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26535
diff changeset
    26
    val thy = Thm.theory_of_thm thm;
28814
463c9e9111ae clarified Thm.proof_body_of vs. Thm.proof_of;
wenzelm
parents: 28800
diff changeset
    27
    val thm' = Reconstruct.reconstruct_proof thy (Thm.prop_of thm) (Thm.proof_of thm);
26626
c6231d64d264 rep_cterm/rep_thm: no longer dereference theory_ref;
wenzelm
parents: 26535
diff changeset
    28
  in Reconstruct.expand_proof thy [("", NONE)] thm' end; (* FIXME *)
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    29
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    30
fun subsets [] = [[]]
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    31
  | subsets (x::xs) =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    32
      let val ys = subsets xs
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    33
      in ys @ map (cons x) ys end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    34
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    35
val pred_of = fst o dest_Const o head_of;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    36
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    37
fun strip_all' used names (Const ("all", _) $ Abs (s, T, t)) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    38
      let val (s', names') = (case names of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    39
          [] => (Name.variant used s, [])
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    40
        | name :: names' => (name, names'))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    41
      in strip_all' (s'::used) names' (subst_bound (Free (s', T), t)) end
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    42
  | strip_all' used names ((t as Const ("==>", _) $ P) $ Q) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    43
      t $ strip_all' used names Q
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    44
  | strip_all' _ _ t = t;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    45
29281
b22ccb3998db eliminated OldTerm.add_term_free_names;
wenzelm
parents: 29271
diff changeset
    46
fun strip_all t = strip_all' (Term.add_free_names t []) [] t;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    47
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    48
fun strip_one name (Const ("all", _) $ Abs (s, T, Const ("==>", _) $ P $ Q)) =
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    49
      (subst_bound (Free (name, T), P), subst_bound (Free (name, T), Q))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    50
  | strip_one _ (Const ("==>", _) $ P $ Q) = (P, Q);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    51
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    52
fun relevant_vars prop = fold (fn ((a, i), T) => fn vs =>
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    53
     (case strip_type T of
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
    54
        (_, Type (s, _)) => if s = @{type_name bool} then (a, T) :: vs else vs
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    55
      | _ => vs)) (Term.add_vars prop []) [];
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    56
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    57
val attach_typeS = map_types (map_atyps
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    58
  (fn TFree (s, []) => TFree (s, HOLogic.typeS)
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    59
    | TVar (ixn, []) => TVar (ixn, HOLogic.typeS)
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    60
    | T => T));
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    61
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    62
fun dt_of_intrs thy vs nparms intrs =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    63
  let
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    64
    val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    65
    val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    66
      (Logic.strip_imp_concl (prop_of (hd intrs))));
33957
e9afca2118d4 normalized uncurry take/drop
haftmann
parents: 33955
diff changeset
    67
    val params = map dest_Var (take nparms ts);
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
    68
    val tname = Binding.name (space_implode "_" (Long_Name.base_name s ^ "T" :: vs));
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
    69
    fun constr_of_intr intr = (Binding.name (Long_Name.base_name (name_of_thm intr)),
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
    70
      map (Logic.unvarifyT_global o snd) (subtract (op =) params (rev (Term.add_vars (prop_of intr) []))) @
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    71
        filter_out (equal Extraction.nullT) (map
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
    72
          (Logic.unvarifyT_global o Extraction.etype_of thy vs []) (prems_of intr)),
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    73
            NoSyn);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    74
  in (map (fn a => "'" ^ a) vs @ map (fst o fst) iTs, tname, NoSyn,
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    75
    map constr_of_intr intrs)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    76
  end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    77
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    78
fun mk_rlz T = Const ("realizes", [T, HOLogic.boolT] ---> HOLogic.boolT);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    79
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    80
(** turn "P" into "%r x. realizes r (P x)" **)
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    81
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    82
fun gen_rvar vs (t as Var ((a, 0), T)) =
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    83
      if body_type T <> HOLogic.boolT then t else
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    84
        let
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    85
          val U = TVar (("'" ^ a, 0), [])
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    86
          val Ts = binder_types T;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    87
          val i = length Ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    88
          val xs = map (pair "x") Ts;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    89
          val u = list_comb (t, map Bound (i - 1 downto 0))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    90
        in 
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
    91
          if member (op =) vs a then
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    92
            list_abs (("r", U) :: xs, mk_rlz U $ Bound i $ u)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    93
          else list_abs (xs, mk_rlz Extraction.nullT $ Extraction.nullt $ u)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    94
        end
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    95
  | gen_rvar _ t = t;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    96
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
    97
fun mk_realizes_eqn n vs nparms intrs =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
    98
  let
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
    99
    val intr = map_types Type.strip_sorts (prop_of (hd intrs));
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   100
    val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl intr);
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   101
    val iTs = rev (Term.add_tvars intr []);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   102
    val Tvs = map TVar iTs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   103
    val (h as Const (s, T), us) = strip_comb concl;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   104
    val params = List.take (us, nparms);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   105
    val elTs = List.drop (binder_types T, nparms);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   106
    val predT = elTs ---> HOLogic.boolT;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   107
    val used = map (fst o fst o dest_Var) params;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   108
    val xs = map (Var o apfst (rpair 0))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   109
      (Name.variant_list used (replicate (length elTs) "x") ~~ elTs);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   110
    val rT = if n then Extraction.nullT
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   111
      else Type (space_implode "_" (s ^ "T" :: vs),
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   112
        map (fn a => TVar (("'" ^ a, 0), [])) vs @ Tvs);
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   113
    val r = if n then Extraction.nullt else Var ((Long_Name.base_name s, 0), rT);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   114
    val S = list_comb (h, params @ xs);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   115
    val rvs = relevant_vars S;
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   116
    val vs' = subtract (op =) vs (map fst rvs);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   117
    val rname = space_implode "_" (s ^ "R" :: vs);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   118
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   119
    fun mk_Tprem n v =
17485
c39871c52977 introduced AList module
haftmann
parents: 17325
diff changeset
   120
      let val T = (the o AList.lookup (op =) rvs) v
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   121
      in (Const ("typeof", T --> Type ("Type", [])) $ Var ((v, 0), T),
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   122
        Extraction.mk_typ (if n then Extraction.nullT
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   123
          else TVar (("'" ^ v, 0), [])))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   124
      end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   125
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   126
    val prems = map (mk_Tprem true) vs' @ map (mk_Tprem false) vs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   127
    val ts = map (gen_rvar vs) params;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   128
    val argTs = map fastype_of ts;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   129
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   130
  in ((prems, (Const ("typeof", HOLogic.boolT --> Type ("Type", [])) $ S,
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   131
       Extraction.mk_typ rT)),
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   132
    (prems, (mk_rlz rT $ r $ S,
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   133
       if n then list_comb (Const (rname, argTs ---> predT), ts @ xs)
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   134
       else list_comb (Const (rname, argTs @ [rT] ---> predT), ts @ [r] @ xs))))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   135
  end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   136
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   137
fun fun_of_prem thy rsets vs params rule ivs intr =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   138
  let
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36043
diff changeset
   139
    val ctxt = ProofContext.init_global thy
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   140
    val args = map (Free o apfst fst o dest_Var) ivs;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   141
    val args' = map (Free o apfst fst)
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   142
      (subtract (op =) params (Term.add_vars (prop_of intr) []));
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   143
    val rule' = strip_all rule;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   144
    val conclT = Extraction.etype_of thy vs [] (Logic.strip_imp_concl rule');
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   145
    val used = map (fst o dest_Free) args;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   146
29271
1d685baea08e moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents: 29265
diff changeset
   147
    val is_rec = exists_Const (fn (c, _) => member (op =) rsets c);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   148
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   149
    fun is_meta (Const ("all", _) $ Abs (s, _, P)) = is_meta P
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   150
      | is_meta (Const ("==>", _) $ _ $ Q) = is_meta Q
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   151
      | is_meta (Const (@{const_name Trueprop}, _) $ t) =
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   152
          (case head_of t of
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   153
            Const (s, _) => can (Inductive.the_inductive ctxt) s
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   154
          | _ => true)
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   155
      | is_meta _ = false;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   156
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   157
    fun fun_of ts rts args used (prem :: prems) =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   158
          let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   159
            val T = Extraction.etype_of thy vs [] prem;
20071
8f3e1ddb50e6 replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents: 19806
diff changeset
   160
            val [x, r] = Name.variant_list used ["x", "r"]
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   161
          in if T = Extraction.nullT
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   162
            then fun_of ts rts args used prems
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   163
            else if is_rec prem then
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   164
              if is_meta prem then
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   165
                let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   166
                  val prem' :: prems' = prems;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   167
                  val U = Extraction.etype_of thy vs [] prem';
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   168
                in if U = Extraction.nullT
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   169
                  then fun_of (Free (x, T) :: ts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   170
                    (Free (r, binder_types T ---> HOLogic.unitT) :: rts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   171
                    (Free (x, T) :: args) (x :: r :: used) prems'
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   172
                  else fun_of (Free (x, T) :: ts) (Free (r, U) :: rts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   173
                    (Free (r, U) :: Free (x, T) :: args) (x :: r :: used) prems'
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   174
                end
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   175
              else (case strip_type T of
37678
0040bafffdef "prod" and "sum" replace "*" and "+" respectively
haftmann
parents: 37236
diff changeset
   176
                  (Ts, Type (@{type_name Product_Type.prod}, [T1, T2])) =>
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   177
                    let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   178
                      val fx = Free (x, Ts ---> T1);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   179
                      val fr = Free (r, Ts ---> T2);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   180
                      val bs = map Bound (length Ts - 1 downto 0);
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   181
                      val t = list_abs (map (pair "z") Ts,
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   182
                        HOLogic.mk_prod (list_comb (fx, bs), list_comb (fr, bs)))
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   183
                    in fun_of (fx :: ts) (fr :: rts) (t::args)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   184
                      (x :: r :: used) prems
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   185
                    end
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   186
                | (Ts, U) => fun_of (Free (x, T) :: ts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   187
                    (Free (r, binder_types T ---> HOLogic.unitT) :: rts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   188
                    (Free (x, T) :: args) (x :: r :: used) prems)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   189
            else fun_of (Free (x, T) :: ts) rts (Free (x, T) :: args)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   190
              (x :: used) prems
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   191
          end
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   192
      | fun_of ts rts args used [] =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   193
          let val xs = rev (rts @ ts)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   194
          in if conclT = Extraction.nullT
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   195
            then list_abs_free (map dest_Free xs, HOLogic.unit)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   196
            else list_abs_free (map dest_Free xs, list_comb
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   197
              (Free ("r" ^ Long_Name.base_name (name_of_thm intr),
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   198
                map fastype_of (rev args) ---> conclT), rev args))
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   199
          end
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   200
13921
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   201
  in fun_of args' [] (rev args) used (Logic.strip_imp_prems rule') end;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   202
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   203
fun indrule_realizer thy induct raw_induct rsets params vs rec_names rss intrs dummies =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   204
  let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   205
    val concls = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct));
31986
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31784
diff changeset
   206
    val premss = map_filter (fn (s, rs) => if member (op =) rsets s then
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31784
diff changeset
   207
      SOME (rs, map (fn (_, r) => nth (prems_of raw_induct)
a68f88d264f7 dropped find_index_eq
haftmann
parents: 31784
diff changeset
   208
        (find_index (fn prp => prp = prop_of r) (map prop_of intrs))) rs) else NONE) rss;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   209
    val fs = maps (fn ((intrs, prems), dummy) =>
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   210
      let
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   211
        val fs = map (fn (rule, (ivs, intr)) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   212
          fun_of_prem thy rsets vs params rule ivs intr) (prems ~~ intrs)
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   213
      in
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   214
        if dummy then Const (@{const_name default},
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   215
            HOLogic.unitT --> body_type (fastype_of (hd fs))) :: fs
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   216
        else fs
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   217
      end) (premss ~~ dummies);
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16123
diff changeset
   218
    val frees = fold Term.add_frees fs [];
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   219
    val Ts = map fastype_of fs;
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   220
    fun name_of_fn intr = "r" ^ Long_Name.base_name (name_of_thm intr)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   221
  in
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   222
    fst (fold_map (fn concl => fn names =>
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   223
      let val T = Extraction.etype_of thy vs [] concl
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   224
      in if T = Extraction.nullT then (Extraction.nullt, names) else
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   225
        let
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   226
          val Type ("fun", [U, _]) = T;
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   227
          val a :: names' = names
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32131
diff changeset
   228
        in (list_abs_free (("x", U) :: map_filter (fn intr =>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   229
          Option.map (pair (name_of_fn intr))
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   230
            (AList.lookup (op =) frees (name_of_fn intr))) intrs,
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   231
          list_comb (Const (a, Ts ---> T), fs) $ Free ("x", U)), names')
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   232
        end
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   233
      end) concls rec_names)
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   234
  end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   235
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   236
fun add_dummy name dname (x as (_, (vs, s, mfx, cs))) =
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   237
  if Binding.eq_name (name, s) then (true, (vs, s, mfx, (dname, [HOLogic.unitT], NoSyn) :: cs))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   238
  else x;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   239
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   240
fun add_dummies f [] _ thy =
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   241
      (([], NONE), thy)
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   242
  | add_dummies f dts used thy =
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   243
      thy
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   244
      |> f (map snd dts)
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   245
      |-> (fn dtinfo => pair (map fst dts, SOME dtinfo))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
   246
    handle Datatype_Aux.Datatype_Empty name' =>
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   247
      let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   248
        val name = Long_Name.base_name name';
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   249
        val dname = Name.variant used "Dummy";
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   250
      in
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   251
        thy
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   252
        |> add_dummies f (map (add_dummy (Binding.name name) (Binding.name dname)) dts) (dname :: used)
14888
99ac3eb0f84e add_dummies no longer uses transform_error but handles specific
berghofe
parents: 13928
diff changeset
   253
      end;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   254
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   255
fun mk_realizer thy vs (name, rule, rrule, rlz, rt) =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   256
  let
13725
12404b452034 Changed format of realizers / correctness proofs.
berghofe
parents: 13710
diff changeset
   257
    val rvs = map fst (relevant_vars (prop_of rule));
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16123
diff changeset
   258
    val xs = rev (Term.add_vars (prop_of rule) []);
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   259
    val vs1 = map Var (filter_out (fn ((a, _), _) => member (op =) rvs a) xs);
16861
7446b4be013b tuned fold on terms;
wenzelm
parents: 16123
diff changeset
   260
    val rlzvs = rev (Term.add_vars (prop_of rrule) []);
17485
c39871c52977 introduced AList module
haftmann
parents: 17325
diff changeset
   261
    val vs2 = map (fn (ixn, _) => Var (ixn, (the o AList.lookup (op =) rlzvs) ixn)) xs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   262
    val rs = map Var (subtract (op = o pairself fst) xs rlzvs);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   263
    val rlz' = fold_rev Logic.all rs (prop_of rrule)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   264
  in (name, (vs,
33338
de76079f973a eliminated some old folds;
wenzelm
parents: 33244
diff changeset
   265
    if rt = Extraction.nullt then rt else fold_rev lambda vs1 rt,
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   266
    Extraction.abs_corr_shyps thy rule vs vs2
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   267
      (ProofRewriteRules.un_hhf_proof rlz' (attach_typeS rlz)
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   268
         (fold_rev Proofterm.forall_intr_proof' rs (prf_of rrule)))))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   269
  end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   270
24157
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   271
fun rename tab = map (fn x => the_default x (AList.lookup op = tab x));
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   272
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   273
fun add_ind_realizer rsets intrs induct raw_induct elims vs thy =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   274
  let
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   275
    val qualifier = Long_Name.qualifier (name_of_thm induct);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   276
    val inducts = Global_Theory.get_thms thy (Long_Name.qualify qualifier "inducts");
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   277
    val iTs = rev (Term.add_tvars (prop_of (hd intrs)) []);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   278
    val ar = length vs + length iTs;
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   279
    val params = Inductive.params_of raw_induct;
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   280
    val arities = Inductive.arities_of raw_induct;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   281
    val nparms = length params;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   282
    val params' = map dest_Var params;
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   283
    val rss = Inductive.partition_rules raw_induct intrs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   284
    val rss' = map (fn (((s, rs), (_, arity)), elim) =>
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   285
      (s, (Inductive.infer_intro_vars elim arity rs ~~ rs)))
22790
e1cff9268177 Moved functions infer_intro_vars, arities_of, params_of, and
berghofe
parents: 22606
diff changeset
   286
        (rss ~~ arities ~~ elims);
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   287
    val (prfx, _) = split_last (Long_Name.explode (fst (hd rss)));
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   288
    val tnames = map (fn s => space_implode "_" (s ^ "T" :: vs)) rsets;
16123
1381e90c2694 Theory.restore_naming;
wenzelm
parents: 15706
diff changeset
   289
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   290
    val thy1 = thy |>
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24157
diff changeset
   291
      Sign.root_path |>
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   292
      Sign.add_path (Long_Name.implode prfx);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   293
    val (ty_eqs, rlz_eqs) = split_list
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   294
      (map (fn (s, rs) => mk_realizes_eqn (not (member (op =) rsets s)) vs nparms rs) rss);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   295
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   296
    val thy1' = thy1 |>
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   297
      Theory.copy |>
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   298
      Sign.add_types (map (fn s => (Binding.name (Long_Name.base_name s), ar, NoSyn)) tnames) |>
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   299
        Extraction.add_typeof_eqns_i ty_eqs;
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   300
    val dts = map_filter (fn (s, rs) => if member (op =) rsets s then
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   301
      SOME (dt_of_intrs thy1' vs nparms rs) else NONE) rss;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   302
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   303
    (** datatype representing computational content of inductive set **)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   304
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
   305
    val ((dummies, some_dt_names), thy2) =
18008
f193815cab2c swapped add_datatype result
haftmann
parents: 17959
diff changeset
   306
      thy1
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 31668
diff changeset
   307
      |> add_dummies (Datatype.add_datatype
32125
10e1a6ea8df9 dropped ancient flat_names option
haftmann
parents: 31986
diff changeset
   308
           { strict = false, quiet = false } (map (Binding.name_of o #2) dts))
18314
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   309
           (map (pair false) dts) []
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   310
      ||> Extraction.add_typeof_eqns_i ty_eqs
4595eb4627fa oriented pairs theory * 'a to 'a * theory
haftmann
parents: 18008
diff changeset
   311
      ||> Extraction.add_realizes_eqns_i rlz_eqs;
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
   312
    val dt_names = these some_dt_names;
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   313
    val case_thms = map (#case_rewrites o Datatype.the_info thy2) dt_names;
31783
cfbe9609ceb1 add_datatypes does not yield particular rules any longer
haftmann
parents: 31781
diff changeset
   314
    val rec_thms = if null dt_names then []
31784
bd3486c57ba3 tuned interfaces of datatype module
haftmann
parents: 31783
diff changeset
   315
      else (#rec_rewrites o Datatype.the_info thy2) (hd dt_names);
19046
bc5c6c9b114e removed distinct, renamed gen_distinct to distinct;
wenzelm
parents: 18929
diff changeset
   316
    val rec_names = distinct (op =) (map (fst o dest_Const o head_of o fst o
31781
861e675f01e6 add_datatype interface yields type names and less rules
haftmann
parents: 31723
diff changeset
   317
      HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) rec_thms);
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 31177
diff changeset
   318
    val (constrss, _) = fold_map (fn (s, rs) => fn (recs, dummies) =>
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 31177
diff changeset
   319
      if member (op =) rsets s then
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   320
        let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   321
          val (d :: dummies') = dummies;
19473
wenzelm
parents: 19046
diff changeset
   322
          val (recs1, recs2) = chop (length rs) (if d then tl recs else recs)
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 31177
diff changeset
   323
        in (map (head_of o hd o rev o snd o strip_comb o fst o
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 31177
diff changeset
   324
          HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) recs1, (recs2, dummies'))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   325
        end
31458
b1cf26f2919b avoid Library.foldl_map
haftmann
parents: 31177
diff changeset
   326
      else (replicate (length rs) Extraction.nullt, (recs, dummies)))
31781
861e675f01e6 add_datatype interface yields type names and less rules
haftmann
parents: 31723
diff changeset
   327
        rss (rec_thms, dummies);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   328
    val rintrs = map (fn (intr, c) => attach_typeS (Envir.eta_contract
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   329
      (Extraction.realizes_of thy2 vs
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   330
        (if c = Extraction.nullt then c else list_comb (c, map Var (rev
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   331
          (subtract (op =) params' (Term.add_vars (prop_of intr) []))))) (prop_of intr))))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32131
diff changeset
   332
            (maps snd rss ~~ flat constrss);
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   333
    val (rlzpreds, rlzpreds') =
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   334
      rintrs |> map (fn rintr =>
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   335
        let
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   336
          val Const (s, T) = head_of (HOLogic.dest_Trueprop (Logic.strip_assums_concl rintr));
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   337
          val s' = Long_Name.base_name s;
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   338
          val T' = Logic.unvarifyT_global T;
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   339
        in (((s', T'), NoSyn), (Const (s, T'), Free (s', T'))) end)
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   340
      |> distinct (op = o pairself (#1 o #1))
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   341
      |> map (apfst (apfst (apfst Binding.name)))
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   342
      |> split_list;
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30280
diff changeset
   343
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   344
    val rlzparams = map (fn Var ((s, _), T) => (s, Logic.unvarifyT_global T))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   345
      (List.take (snd (strip_comb
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   346
        (HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd rintrs)))), nparms));
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   347
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   348
    (** realizability predicate **)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   349
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   350
    val (ind_info, thy3') = thy2 |>
33726
0878aecbf119 eliminated slightly odd name space grouping -- now managed by Isar toplevel;
wenzelm
parents: 33671
diff changeset
   351
      Inductive.add_inductive_global
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   352
        {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false,
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33666
diff changeset
   353
          no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   354
        rlzpreds rlzparams (map (fn (rintr, intr) =>
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   355
          ((Binding.name (Long_Name.base_name (name_of_thm intr)), []),
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   356
           subst_atomic rlzpreds' (Logic.unvarify_global rintr)))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   357
             (rintrs ~~ maps snd rss)) [] ||>
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   358
      Sign.root_path;
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   359
    val thy3 = fold (Global_Theory.hide_fact false o name_of_thm) (#intrs ind_info) thy3';
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   360
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   361
    (** realizer for induction rule **)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   362
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   363
    val Ps = map_filter (fn _ $ M $ P => if member (op =) rsets (pred_of M) then
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   364
      SOME (fst (fst (dest_Var (head_of P)))) else NONE)
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   365
        (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of raw_induct)));
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   366
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   367
    fun add_ind_realizer Ps thy =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   368
      let
24157
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   369
        val vs' = rename (map (pairself (fst o fst o dest_Var))
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   370
          (params ~~ List.take (snd (strip_comb (HOLogic.dest_Trueprop
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   371
            (hd (prems_of (hd inducts))))), nparms))) vs;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   372
        val rs = indrule_realizer thy induct raw_induct rsets params'
24157
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   373
          (vs' @ Ps) rec_names rss' intrs dummies;
409cd6eaa7ea Added renaming function to prevent correctness proof for realizer
berghofe
parents: 23590
diff changeset
   374
        val rlzs = map (fn (r, ind) => Extraction.realizes_of thy (vs' @ Ps) r
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   375
          (prop_of ind)) (rs ~~ inducts);
29281
b22ccb3998db eliminated OldTerm.add_term_free_names;
wenzelm
parents: 29271
diff changeset
   376
        val used = fold Term.add_free_names rlzs [];
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   377
        val rnames = Name.variant_list used (replicate (length inducts) "r");
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   378
        val rnames' = Name.variant_list
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   379
          (used @ rnames) (replicate (length intrs) "s");
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   380
        val rlzs' as (prems, _, _) :: _ = map (fn (rlz, name) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   381
          let
35845
e5980f0ad025 renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents: 35625
diff changeset
   382
            val (P, Q) = strip_one name (Logic.unvarify_global rlz);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   383
            val Q' = strip_all' [] rnames' Q
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   384
          in
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   385
            (Logic.strip_imp_prems Q', P, Logic.strip_imp_concl Q')
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   386
          end) (rlzs ~~ rnames);
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   387
        val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   388
          (fn (_, _ $ P, _ $ Q) => HOLogic.mk_imp (P, Q)) rlzs'));
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36945
diff changeset
   389
        val rews = map mk_meta_eq (@{thm fst_conv} :: @{thm snd_conv} :: rec_thms);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   390
        val thm = Goal.prove_global thy []
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   391
          (map attach_typeS prems) (attach_typeS concl)
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   392
          (fn {prems, ...} => EVERY
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   393
          [rtac (#raw_induct ind_info) 1,
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   394
           rewrite_goals_tac rews,
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   395
           REPEAT ((resolve_tac prems THEN_ALL_NEW EVERY'
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35402
diff changeset
   396
             [K (rewrite_goals_tac rews), Object_Logic.atomize_prems_tac,
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   397
              DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE]]) 1)]);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   398
        val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   399
          (Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   400
        val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp)))
33968
f94fb13ecbb3 modernized structures and tuned headers of datatype package modules; joined former datatype.ML and datatype_rep_proofs.ML
haftmann
parents: 33957
diff changeset
   401
          (Datatype_Aux.split_conj_thm thm');
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   402
        val ([thms'], thy'') = Global_Theory.add_thmss
30435
e62d6ecab6ad explicit Binding.qualified_name -- prevents implicitly qualified bstring;
wenzelm
parents: 30364
diff changeset
   403
          [((Binding.qualified_name (space_implode "_"
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30345
diff changeset
   404
             (Long_Name.qualify qualifier "inducts" :: vs' @ Ps @
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29389
diff changeset
   405
               ["correctness"])), thms), [])] thy';
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   406
        val realizers = inducts ~~ thms' ~~ rlzs ~~ rs;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   407
      in
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   408
        Extraction.add_realizers_i
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   409
          (map (fn (((ind, corr), rlz), r) =>
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   410
              mk_realizer thy'' (vs' @ Ps) (Thm.derivation_name ind, ind, corr, rlz, r))
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   411
            realizers @ (case realizers of
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   412
             [(((ind, corr), rlz), r)] =>
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   413
               [mk_realizer thy'' (vs' @ Ps) (Long_Name.qualify qualifier "induct",
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   414
                  ind, corr, rlz, r)]
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   415
           | _ => [])) thy''
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   416
      end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   417
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   418
    (** realizer for elimination rules **)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   419
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   420
    val case_names = map (fst o dest_Const o head_of o fst o HOLogic.dest_eq o
31781
861e675f01e6 add_datatype interface yields type names and less rules
haftmann
parents: 31723
diff changeset
   421
      HOLogic.dest_Trueprop o prop_of o hd) case_thms;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   422
13921
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   423
    fun add_elim_realizer Ps
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   424
      (((((elim, elimR), intrs), case_thms), case_name), dummy) thy =
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   425
      let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   426
        val (prem :: prems) = prems_of elim;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   427
        fun reorder1 (p, (_, intr)) =
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   428
          fold (fn ((s, _), T) => Logic.all (Free (s, T)))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   429
            (subtract (op =) params' (Term.add_vars (prop_of intr) []))
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   430
            (strip_all p);
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   431
        fun reorder2 ((ivs, intr), i) =
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   432
          let val fs = subtract (op =) params' (Term.add_vars (prop_of intr) [])
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   433
          in fold (lambda o Var) fs (list_comb (Bound (i + length ivs), ivs)) end;
13921
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   434
        val p = Logic.list_implies
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   435
          (map reorder1 (prems ~~ intrs) @ [prem], concl_of elim);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   436
        val T' = Extraction.etype_of thy (vs @ Ps) [] p;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   437
        val T = if dummy then (HOLogic.unitT --> body_type T') --> T' else T';
13921
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   438
        val Ts = map (Extraction.etype_of thy (vs @ Ps) []) (prems_of elim);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   439
        val r = if null Ps then Extraction.nullt
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   440
          else list_abs (map (pair "x") Ts, list_comb (Const (case_name, T),
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   441
            (if dummy then
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 33968
diff changeset
   442
               [Abs ("x", HOLogic.unitT, Const (@{const_name default}, body_type T))]
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   443
             else []) @
13921
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   444
            map reorder2 (intrs ~~ (length prems - 1 downto 0)) @
69c627b6b28d Fixed problem in add_elim_realizer which caused bound variables to
berghofe
parents: 13725
diff changeset
   445
            [Bound (length prems)]));
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   446
        val rlz = Extraction.realizes_of thy (vs @ Ps) r (prop_of elim);
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   447
        val rlz' = attach_typeS (strip_all (Logic.unvarify_global rlz));
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   448
        val rews = map mk_meta_eq case_thms;
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   449
        val thm = Goal.prove_global thy []
26711
3a478bfa1650 prove_global: pass context;
wenzelm
parents: 26663
diff changeset
   450
          (Logic.strip_imp_prems rlz') (Logic.strip_imp_concl rlz') (fn {prems, ...} => EVERY
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   451
          [cut_facts_tac [hd prems] 1,
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   452
           etac elimR 1,
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   453
           ALLGOALS (asm_simp_tac HOL_basic_ss),
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   454
           rewrite_goals_tac rews,
35625
9c818cab0dd0 modernized structure Object_Logic;
wenzelm
parents: 35402
diff changeset
   455
           REPEAT ((resolve_tac prems THEN_ALL_NEW (Object_Logic.atomize_prems_tac THEN'
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   456
             DEPTH_SOLVE_1 o FIRST' [atac, etac allE, etac impE])) 1)]);
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37678
diff changeset
   457
        val (thm', thy') = Global_Theory.store_thm (Binding.qualified_name (space_implode "_"
29579
cb520b766e00 binding replaces bstring
haftmann
parents: 29389
diff changeset
   458
          (name_of_thm elim :: vs @ Ps @ ["correctness"])), thm) thy
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   459
      in
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   460
        Extraction.add_realizers_i
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   461
          [mk_realizer thy' (vs @ Ps) (name_of_thm elim, elim, thm', rlz, r)] thy'
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   462
      end;
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   463
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   464
    (** add realizers to theory **)
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   465
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   466
    val thy4 = fold add_ind_realizer (subsets Ps) thy3;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   467
    val thy5 = Extraction.add_realizers_i
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   468
      (map (mk_realizer thy4 vs) (map (fn (((rule, rrule), rlz), c) =>
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   469
         (name_of_thm rule, rule, rrule, rlz,
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   470
            list_comb (c, map Var (subtract (op =) params' (rev (Term.add_vars (prop_of rule) []))))))
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32131
diff changeset
   471
              (maps snd rss ~~ #intrs ind_info ~~ rintrs ~~ flat constrss))) thy4;
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32131
diff changeset
   472
    val elimps = map_filter (fn ((s, intrs), p) =>
36692
54b64d4ad524 farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents: 36610
diff changeset
   473
      if member (op =) rsets s then SOME (p, intrs) else NONE)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   474
        (rss' ~~ (elims ~~ #elims ind_info));
33244
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   475
    val thy6 =
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   476
      fold (fn p as (((((elim, _), _), _), _), _) =>
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   477
        add_elim_realizer [] p #>
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   478
        add_elim_realizer [fst (fst (dest_Var (HOLogic.dest_Trueprop (concl_of elim))))] p)
db230399f890 Datatype.read_typ: standard argument order;
wenzelm
parents: 33171
diff changeset
   479
      (elimps ~~ case_thms ~~ case_names ~~ dummies) thy5;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   480
24712
64ed05609568 proper Sign operations instead of Theory aliases;
wenzelm
parents: 24157
diff changeset
   481
  in Sign.restore_naming thy thy6 end;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   482
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   483
fun add_ind_realizers name rsets thy =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   484
  let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   485
    val (_, {intrs, induct, raw_induct, elims, ...}) =
36610
bafd82950e24 renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
wenzelm
parents: 36043
diff changeset
   486
      Inductive.the_inductive (ProofContext.init_global thy) name;
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   487
    val vss = sort (int_ord o pairself length)
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   488
      (subsets (map fst (relevant_vars (concl_of (hd intrs)))))
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   489
  in
37233
b78f31ca4675 Adapted to new format of proof terms containing explicit proofs of class membership.
berghofe
parents: 36945
diff changeset
   490
    fold_rev (add_ind_realizer rsets intrs induct raw_induct elims) vss thy
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   491
  end
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   492
20897
3f8d2834b2c4 attribute: Context.mapping;
wenzelm
parents: 20071
diff changeset
   493
fun rlz_attrib arg = Thm.declaration_attribute (fn thm => Context.mapping
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   494
  let
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   495
    fun err () = error "ind_realizer: bad rule";
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   496
    val sets =
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   497
      (case HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of thm)) of
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   498
           [_] => [pred_of (HOLogic.dest_Trueprop (hd (prems_of thm)))]
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 21858
diff changeset
   499
         | xs => map (pred_of o fst o HOLogic.dest_imp) xs)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   500
         handle TERM _ => err () | Empty => err ();
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   501
  in 
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   502
    add_ind_realizers (hd sets)
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   503
      (case arg of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 14981
diff changeset
   504
        NONE => sets | SOME NONE => []
33040
cffdb7b28498 removed old-style \ and \\ infixes
haftmann
parents: 32952
diff changeset
   505
      | SOME (SOME sets') => subtract (op =) sets' sets)
20897
3f8d2834b2c4 attribute: Context.mapping;
wenzelm
parents: 20071
diff changeset
   506
  end I);
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   507
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18358
diff changeset
   508
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
   509
  Attrib.setup @{binding ind_realizer}
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
   510
    ((Scan.option (Scan.lift (Args.$$$ "irrelevant") |--
35402
115a5a95710a clarified @{const_name} vs. @{const_abbrev};
wenzelm
parents: 35364
diff changeset
   511
      Scan.option (Scan.lift (Args.colon) |-- Scan.repeat1 (Args.const true)))) >> rlz_attrib)
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
   512
    "add realizers for inductive set";
13710
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   513
75bec2c1bfd5 New package for constructing realizers for introduction and elimination
berghofe
parents:
diff changeset
   514
end;
15706
bc264e730103 *** MESSAGE REFERS TO PREVIOUS VERSION ***
wenzelm
parents: 15703
diff changeset
   515