src/HOL/Tools/inductive_set.ML
author boehmes
Mon, 08 Nov 2010 12:13:44 +0100
changeset 40424 7550b2cba1cb
parent 38864 4abe644fcea5
child 41472 f6ab14e61604
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:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
     1
(*  Title:      HOL/Tools/inductive_set.ML
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     3
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     4
Wrapper for defining inductive sets using package for inductive predicates,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     5
including infrastructure for converting between predicates and sets.
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     6
*)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     7
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
     8
signature INDUCTIVE_SET =
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
     9
sig
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    10
  val to_set_att: thm list -> attribute
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    11
  val to_pred_att: thm list -> attribute
32306
19f55947d4d5 removed debug messages; exported to_pred in InductiveSet; added further display function; adjusted mode analysis
bulwahn
parents: 32287
diff changeset
    12
  val to_pred : thm list -> Context.generic -> thm -> thm
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    13
  val pred_set_conv_att: attribute
24815
f7093e90f36c tuned internal interfaces: flags record, added kind for results;
wenzelm
parents: 24745
diff changeset
    14
  val add_inductive_i:
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
    15
    Inductive.inductive_flags ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29389
diff changeset
    16
    ((binding * typ) * mixfix) list ->
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    17
    (string * typ) list ->
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    18
    (Attrib.binding * term) list -> thm list ->
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
    19
    local_theory -> Inductive.inductive_result * local_theory
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    20
  val add_inductive: bool -> bool ->
29581
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29389
diff changeset
    21
    (binding * string option * mixfix) list ->
b3b33e0298eb binding is alias for Binding.T
haftmann
parents: 29389
diff changeset
    22
    (binding * string option * mixfix) list ->
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    23
    (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
    24
    bool -> local_theory -> Inductive.inductive_result * local_theory
28723
c4fcffe0fe48 exported codegen_preproc
haftmann
parents: 28084
diff changeset
    25
  val codegen_preproc: theory -> thm list -> thm list
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    26
  val setup: theory -> theory
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    27
end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    28
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
    29
structure Inductive_Set: INDUCTIVE_SET =
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    30
struct
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    31
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    32
(**** simplify {(x1, ..., xn). (x1, ..., xn) : S} to S ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    33
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    34
val collect_mem_simproc =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38665
diff changeset
    35
  Simplifier.simproc_global @{theory Set} "Collect_mem" ["Collect t"] (fn thy => fn ss =>
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 34986
diff changeset
    36
    fn S as Const (@{const_name Collect}, Type ("fun", [_, T])) $ t =>
34903
d75704c60768 dropped unused binding
haftmann
parents: 33766
diff changeset
    37
         let val (u, _, ps) = HOLogic.strip_psplits t
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    38
         in case u of
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
    39
           (c as Const (@{const_name Set.member}, _)) $ q $ S' =>
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
    40
             (case try (HOLogic.strip_ptuple ps) q of
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    41
                NONE => NONE
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    42
              | SOME ts =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    43
                  if not (loose_bvar (S', 0)) andalso
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    44
                    ts = map Bound (length ps downto 0)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    45
                  then
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    46
                    let val simp = full_simp_tac (Simplifier.inherit_context ss
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
    47
                      (HOL_basic_ss addsimps [@{thm split_paired_all}, @{thm split_conv}])) 1
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    48
                    in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    49
                      SOME (Goal.prove (Simplifier.the_context ss) [] []
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    50
                        (Const ("==", T --> T --> propT) $ S $ S')
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    51
                        (K (EVERY
24815
f7093e90f36c tuned internal interfaces: flags record, added kind for results;
wenzelm
parents: 24745
diff changeset
    52
                          [rtac eq_reflection 1, rtac @{thm subset_antisym} 1,
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    53
                           rtac subsetI 1, dtac CollectD 1, simp,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    54
                           rtac subsetI 1, rtac CollectI 1, simp])))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    55
                    end
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    56
                  else NONE)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    57
         | _ => NONE
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    58
         end
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    59
     | _ => NONE);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    60
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    61
(***********************************************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    62
(* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    63
(* and        (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y})  *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    64
(* used for converting "strong" (co)induction rules                                *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    65
(***********************************************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    66
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    67
val anyt = Free ("t", TFree ("'t", []));
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    68
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    69
fun strong_ind_simproc tab =
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38665
diff changeset
    70
  Simplifier.simproc_global_i @{theory HOL} "strong_ind" [anyt] (fn thy => fn ss => fn t =>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    71
    let
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    72
      fun close p t f =
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    73
        let val vs = Term.add_vars t []
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    74
        in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
27330
1af2598b5f7d Logic.all/mk_equals/mk_implies;
wenzelm
parents: 26988
diff changeset
    75
          (p (fold (Logic.all o Var) vs t) f)
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    76
        end;
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38715
diff changeset
    77
      fun mkop @{const_name HOL.conj} T x =
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 34986
diff changeset
    78
            SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x)
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38715
diff changeset
    79
        | mkop @{const_name HOL.disj} T x =
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 34986
diff changeset
    80
            SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x)
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    81
        | mkop _ _ _ = NONE;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    82
      fun mk_collect p T t =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    83
        let val U = HOLogic.dest_setT T
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    84
        in HOLogic.Collect_const U $
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
    85
          HOLogic.mk_psplits (HOLogic.flat_tuple_paths p) U HOLogic.boolT t
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    86
        end;
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
    87
      fun decomp (Const (s, _) $ ((m as Const (@{const_name Set.member},
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    88
            Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    89
              mkop s T (m, p, S, mk_collect p T (head_of u))
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
    90
        | decomp (Const (s, _) $ u $ ((m as Const (@{const_name Set.member},
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    91
            Type (_, [_, Type (_, [T, _])]))) $ p $ S)) =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    92
              mkop s T (m, p, mk_collect p T (head_of u), S)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    93
        | decomp _ = NONE;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
    94
      val simp = full_simp_tac (Simplifier.inherit_context ss
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
    95
        (HOL_basic_ss addsimps [mem_Collect_eq, @{thm split_conv}])) 1;
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    96
      fun mk_rew t = (case strip_abs_vars t of
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    97
          [] => NONE
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    98
        | xs => (case decomp (strip_abs_body t) of
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
    99
            NONE => NONE
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   100
          | SOME (bop, (m, p, S, S')) =>
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   101
              SOME (close (Goal.prove (Simplifier.the_context ss) [] [])
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   102
                (Logic.mk_equals (t, list_abs (xs, m $ p $ (bop $ S $ S'))))
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   103
                (K (EVERY
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   104
                  [rtac eq_reflection 1, REPEAT (rtac ext 1), rtac iffI 1,
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   105
                   EVERY [etac conjE 1, rtac IntI 1, simp, simp,
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   106
                     etac IntE 1, rtac conjI 1, simp, simp] ORELSE
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   107
                   EVERY [etac disjE 1, rtac UnI1 1, simp, rtac UnI2 1, simp,
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   108
                     etac UnE 1, rtac disjI1 1, simp, rtac disjI2 1, simp]])))
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   109
                handle ERROR _ => NONE))
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   110
    in
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   111
      case strip_comb t of
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   112
        (h as Const (name, _), ts) => (case Symtab.lookup tab name of
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   113
          SOME _ =>
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   114
            let val rews = map mk_rew ts
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   115
            in
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   116
              if forall is_none rews then NONE
36945
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   117
              else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   118
                (map2 (fn SOME r => K r | NONE => Thm.reflexive o cterm_of thy)
9bec62c10714 less pervasive names from structure Thm;
wenzelm
parents: 36692
diff changeset
   119
                   rews ts) (Thm.reflexive (cterm_of thy h)))
23849
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   120
            end
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   121
        | NONE => NONE)
2a0e24c74593 strong_ind_simproc now only rewrites arguments of inductive predicates.
berghofe
parents: 23764
diff changeset
   122
      | _ => NONE
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   123
    end);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   124
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   125
(* only eta contract terms occurring as arguments of functions satisfying p *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   126
fun eta_contract p =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   127
  let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   128
    fun eta b (Abs (a, T, body)) =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   129
          (case eta b body of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   130
             body' as (f $ Bound 0) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   131
               if loose_bvar1 (f, 0) orelse not b then Abs (a, T, body')
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   132
               else incr_boundvars ~1 f
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   133
           | body' => Abs (a, T, body'))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   134
      | eta b (t $ u) = eta b t $ eta (p (head_of t)) u
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   135
      | eta b t = t
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   136
  in eta false end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   137
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   138
fun eta_contract_thm p =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   139
  Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   140
    Thm.transitive (Thm.eta_conversion ct)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   141
      (Thm.symmetric (Thm.eta_conversion
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   142
        (cterm_of (theory_of_cterm ct) (eta_contract p (term_of ct)))))));
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   143
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   144
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   145
(***********************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   146
(* rules for converting between predicate and set notation *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   147
(*                                                         *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   148
(* rules for converting predicates to sets have the form   *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   149
(* P (%x y. (x, y) : s) = (%x y. (x, y) : S s)             *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   150
(*                                                         *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   151
(* rules for converting sets to predicates have the form   *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   152
(* S {(x, y). p x y} = {(x, y). P p x y}                   *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   153
(*                                                         *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   154
(* where s and p are parameters                            *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   155
(***********************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   156
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33459
diff changeset
   157
structure PredSetConvData = Generic_Data
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   158
(
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   159
  type T =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   160
    {(* rules for converting predicates to sets *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   161
     to_set_simps: thm list,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   162
     (* rules for converting sets to predicates *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   163
     to_pred_simps: thm list,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   164
     (* arities of functions of type t set => ... => u set *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   165
     set_arities: (typ * (int list list option list * int list list option)) list Symtab.table,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   166
     (* arities of functions of type (t => ... => bool) => u => ... => bool *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   167
     pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table};
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   168
  val empty = {to_set_simps = [], to_pred_simps = [],
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   169
    set_arities = Symtab.empty, pred_arities = Symtab.empty};
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   170
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 33459
diff changeset
   171
  fun merge
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   172
    ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   173
      set_arities = set_arities1, pred_arities = pred_arities1},
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   174
     {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2,
29288
253bcf2a5854 avoid polymorphic equality;
wenzelm
parents: 29064
diff changeset
   175
      set_arities = set_arities2, pred_arities = pred_arities2}) : T =
24039
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23849
diff changeset
   176
    {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2),
273698405054 renamed Drule.add/del/merge_rules to Thm.add/del/merge_thms;
wenzelm
parents: 23849
diff changeset
   177
     to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2),
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   178
     set_arities = Symtab.merge_list op = (set_arities1, set_arities2),
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   179
     pred_arities = Symtab.merge_list op = (pred_arities1, pred_arities2)};
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   180
);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   181
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   182
fun name_type_of (Free p) = SOME p
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   183
  | name_type_of (Const p) = SOME p
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   184
  | name_type_of _ = NONE;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   185
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   186
fun map_type f (Free (s, T)) = Free (s, f T)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   187
  | map_type f (Var (ixn, T)) = Var (ixn, f T)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   188
  | map_type f _ = error "map_type";
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   189
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   190
fun find_most_specific is_inst f eq xs T =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   191
  find_first (fn U => is_inst (T, f U)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   192
    andalso forall (fn U' => eq (f U, f U') orelse not
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   193
      (is_inst (T, f U') andalso is_inst (f U', f U)))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   194
        xs) xs;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   195
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   196
fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   197
    NONE => NONE
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   198
  | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   199
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   200
fun lookup_rule thy f rules = find_most_specific
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   201
  (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   202
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   203
fun infer_arities thy arities (optf, t) fs = case strip_comb t of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   204
    (Abs (s, T, u), []) => infer_arities thy arities (NONE, u) fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   205
  | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   206
  | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   207
      SOME (SOME (_, (arity, _))) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   208
        (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   209
           handle Subscript => error "infer_arities: bad term")
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   210
    | _ => fold (infer_arities thy arities) (map (pair NONE) ts)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   211
      (case optf of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   212
         NONE => fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   213
       | SOME f => AList.update op = (u, the_default f
33049
c38f02fdf35d curried inter as canonical list operation (beware of argument order)
haftmann
parents: 33038
diff changeset
   214
           (Option.map (fn g => inter (op =) g f) (AList.lookup op = fs u))) fs));
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   215
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   216
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   217
(**************************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   218
(*    derive the to_pred equation from the to_set equation    *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   219
(*                                                            *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   220
(* 1. instantiate each set parameter with {(x, y). p x y}     *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   221
(* 2. apply %P. {(x, y). P x y} to both sides of the equation *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   222
(* 3. simplify                                                *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   223
(**************************************************************)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   224
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   225
fun mk_to_pred_inst thy fs =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   226
  map (fn (x, ps) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   227
    let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   228
      val U = HOLogic.dest_setT (fastype_of x);
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   229
      val x' = map_type (K (HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   230
    in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   231
      (cterm_of thy x,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   232
       cterm_of thy (HOLogic.Collect_const U $
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   233
         HOLogic.mk_psplits ps U HOLogic.boolT x'))
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   234
    end) fs;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   235
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   236
fun mk_to_pred_eq p fs optfs' T thm =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   237
  let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   238
    val thy = theory_of_thm thm;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   239
    val insts = mk_to_pred_inst thy fs;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   240
    val thm' = Thm.instantiate ([], insts) thm;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   241
    val thm'' = (case optfs' of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   242
        NONE => thm' RS sym
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   243
      | SOME fs' =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   244
          let
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26534
diff changeset
   245
            val (_, U) = split_last (binder_types T);
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   246
            val Ts = HOLogic.strip_ptupleT fs' U;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   247
            (* FIXME: should cterm_instantiate increment indexes? *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   248
            val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   249
            val (arg_cong_f, _) = arg_cong' |> cprop_of |> Drule.strip_imp_concl |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   250
              Thm.dest_comb |> snd |> Drule.strip_comb |> snd |> hd |> Thm.dest_comb
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   251
          in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   252
            thm' RS (Drule.cterm_instantiate [(arg_cong_f,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   253
              cterm_of thy (Abs ("P", Ts ---> HOLogic.boolT,
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   254
                HOLogic.Collect_const U $ HOLogic.mk_psplits fs' U
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   255
                  HOLogic.boolT (Bound 0))))] arg_cong' RS sym)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   256
          end)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   257
  in
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
   258
    Simplifier.simplify (HOL_basic_ss addsimps [mem_Collect_eq, @{thm split_conv}]
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   259
      addsimprocs [collect_mem_simproc]) thm'' |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   260
        zero_var_indexes |> eta_contract_thm (equal p)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   261
  end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   262
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   263
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   264
(**** declare rules for converting predicates to sets ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   265
26047
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   266
fun add ctxt thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) =
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   267
  case prop_of thm of
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   268
    Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, Type (_, [T, _])) $ lhs $ rhs) =>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   269
      (case body_type T of
35364
b8c62d60195c more antiquotations;
wenzelm
parents: 34986
diff changeset
   270
         @{typ bool} =>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   271
           let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   272
             val thy = Context.theory_of ctxt;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   273
             fun factors_of t fs = case strip_abs_body t of
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
   274
                 Const (@{const_name Set.member}, _) $ u $ S =>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   275
                   if is_Free S orelse is_Var S then
32287
65d5c5b30747 cleaned up abstract tuple operations and named them consistently
haftmann
parents: 32135
diff changeset
   276
                     let val ps = HOLogic.flat_tuple_paths u
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   277
                     in (SOME ps, (S, ps) :: fs) end
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   278
                   else (NONE, fs)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   279
               | _ => (NONE, fs);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   280
             val (h, ts) = strip_comb lhs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   281
             val (pfs, fs) = fold_map factors_of ts [];
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   282
             val ((h', ts'), fs') = (case rhs of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   283
                 Abs _ => (case strip_abs_body rhs of
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
   284
                     Const (@{const_name Set.member}, _) $ u $ S =>
32287
65d5c5b30747 cleaned up abstract tuple operations and named them consistently
haftmann
parents: 32135
diff changeset
   285
                       (strip_comb S, SOME (HOLogic.flat_tuple_paths u))
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   286
                   | _ => error "member symbol on right-hand side expected")
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   287
               | _ => (strip_comb rhs, NONE))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   288
           in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   289
             case (name_type_of h, name_type_of h') of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   290
               (SOME (s, T), SOME (s', T')) =>
26047
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   291
                 if exists (fn (U, _) =>
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   292
                   Sign.typ_instance thy (T', U) andalso
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   293
                   Sign.typ_instance thy (U, T'))
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   294
                     (Symtab.lookup_list set_arities s')
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   295
                 then
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   296
                   (warning ("Ignoring conversion rule for operator " ^ s'); tab)
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   297
                 else
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   298
                   {to_set_simps = thm :: to_set_simps,
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   299
                    to_pred_simps =
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   300
                      mk_to_pred_eq h fs fs' T' thm :: to_pred_simps,
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   301
                    set_arities = Symtab.insert_list op = (s',
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   302
                      (T', (map (AList.lookup op = fs) ts', fs'))) set_arities,
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   303
                    pred_arities = Symtab.insert_list op = (s,
d27b89c95b29 Instead of raising an exception, pred_set_conv now ignores conversion
berghofe
parents: 25978
diff changeset
   304
                      (T, (pfs, fs'))) pred_arities}
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   305
             | _ => error "set / predicate constant expected"
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   306
           end
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   307
       | _ => error "equation between predicates expected")
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   308
  | _ => error "equation expected";
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   309
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   310
val pred_set_conv_att = Thm.declaration_attribute
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   311
  (fn thm => fn ctxt => PredSetConvData.map (add ctxt thm) ctxt);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   312
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   313
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   314
(**** convert theorem in set notation to predicate notation ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   315
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   316
fun is_pred tab t =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   317
  case Option.map (Symtab.lookup tab o fst) (name_type_of t) of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   318
    SOME (SOME _) => true | _ => false;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   319
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   320
fun to_pred_simproc rules =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   321
  let val rules' = map mk_meta_eq rules
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   322
  in
38715
6513ea67d95d renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents: 38665
diff changeset
   323
    Simplifier.simproc_global_i @{theory HOL} "to_pred" [anyt]
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   324
      (fn thy => K (lookup_rule thy (prop_of #> Logic.dest_equals) rules'))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   325
  end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   326
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   327
fun to_pred_proc thy rules t = case lookup_rule thy I rules t of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   328
    NONE => NONE
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   329
  | SOME (lhs, rhs) =>
32035
8e77b6a250d5 tuned/modernized Envir.subst_XXX;
wenzelm
parents: 31998
diff changeset
   330
      SOME (Envir.subst_term
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   331
        (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   332
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   333
fun to_pred thms ctxt thm =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   334
  let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   335
    val thy = Context.theory_of ctxt;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   336
    val {to_pred_simps, set_arities, pred_arities, ...} =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   337
      fold (add ctxt) thms (PredSetConvData.get ctxt);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   338
    val fs = filter (is_Var o fst)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   339
      (infer_arities thy set_arities (NONE, prop_of thm) []);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   340
    (* instantiate each set parameter with {(x, y). p x y} *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   341
    val insts = mk_to_pred_inst thy fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   342
  in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   343
    thm |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   344
    Thm.instantiate ([], insts) |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   345
    Simplifier.full_simplify (HOL_basic_ss addsimprocs
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
   346
      [to_pred_simproc (mem_Collect_eq :: @{thm split_conv} :: to_pred_simps)]) |>
25416
1d8ebaf5f211 to_pred and to_set now save induction and case rule tags.
berghofe
parents: 25016
diff changeset
   347
    eta_contract_thm (is_pred pred_arities) |>
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33278
diff changeset
   348
    Rule_Cases.save thm
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   349
  end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   350
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   351
val to_pred_att = Thm.rule_attribute o to_pred;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   352
    
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   353
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   354
(**** convert theorem in predicate notation to set notation ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   355
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   356
fun to_set thms ctxt thm =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   357
  let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   358
    val thy = Context.theory_of ctxt;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   359
    val {to_set_simps, pred_arities, ...} =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   360
      fold (add ctxt) thms (PredSetConvData.get ctxt);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   361
    val fs = filter (is_Var o fst)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   362
      (infer_arities thy pred_arities (NONE, prop_of thm) []);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   363
    (* instantiate each predicate parameter with %x y. (x, y) : s *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   364
    val insts = map (fn (x, ps) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   365
      let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   366
        val Ts = binder_types (fastype_of x);
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   367
        val T = HOLogic.mk_ptupleT ps Ts;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   368
        val x' = map_type (K (HOLogic.mk_setT T)) x
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   369
      in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   370
        (cterm_of thy x,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   371
         cterm_of thy (list_abs (map (pair "x") Ts, HOLogic.mk_mem
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   372
           (HOLogic.mk_ptuple ps T (map Bound (length ps downto 0)), x'))))
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   373
      end) fs
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   374
  in
25416
1d8ebaf5f211 to_pred and to_set now save induction and case rule tags.
berghofe
parents: 25016
diff changeset
   375
    thm |>
1d8ebaf5f211 to_pred and to_set now save induction and case rule tags.
berghofe
parents: 25016
diff changeset
   376
    Thm.instantiate ([], insts) |>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   377
    Simplifier.full_simplify (HOL_basic_ss addsimps to_set_simps
25487
d96d5808d926 to_set now applies collect_mem_simproc as well.
berghofe
parents: 25416
diff changeset
   378
        addsimprocs [strong_ind_simproc pred_arities, collect_mem_simproc]) |>
33368
b1cf34f1855c modernized structure Rule_Cases;
wenzelm
parents: 33278
diff changeset
   379
    Rule_Cases.save thm
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   380
  end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   381
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   382
val to_set_att = Thm.rule_attribute o to_set;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   383
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   384
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   385
(**** preprocessor for code generator ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   386
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   387
fun codegen_preproc thy =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   388
  let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   389
    val {to_pred_simps, set_arities, pred_arities, ...} =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   390
      PredSetConvData.get (Context.Theory thy);
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   391
    fun preproc thm =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   392
      if exists_Const (fn (s, _) => case Symtab.lookup set_arities s of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   393
          NONE => false
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   394
        | SOME arities => exists (fn (_, (xs, _)) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   395
            forall is_none xs) arities) (prop_of thm)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   396
      then
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   397
        thm |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   398
        Simplifier.full_simplify (HOL_basic_ss addsimprocs
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
   399
          [to_pred_simproc (mem_Collect_eq :: @{thm split_conv} :: to_pred_simps)]) |>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   400
        eta_contract_thm (is_pred pred_arities)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   401
      else thm
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   402
  in map preproc end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   403
37390
8781d80026fc moved inductive_codegen to place where product type is available; tuned structure name
haftmann
parents: 37136
diff changeset
   404
fun code_ind_att optmod = to_pred_att [] #> Inductive_Codegen.add optmod NONE;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   405
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   406
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   407
(**** definition of inductive sets ****)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   408
29389
0a49f940d729 inductive: added fork_mono flag;
wenzelm
parents: 29288
diff changeset
   409
fun add_ind_set_def
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   410
    {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   411
    cs intros monos params cnames_syn lthy =
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   412
  let
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   413
    val thy = ProofContext.theory_of lthy;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   414
    val {set_arities, pred_arities, to_pred_simps, ...} =
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   415
      PredSetConvData.get (Context.Proof lthy);
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   416
    fun infer (Abs (_, _, t)) = infer t
37677
c5a8b612e571 qualified constants Set.member and Set.Collect
haftmann
parents: 37390
diff changeset
   417
      | infer (Const (@{const_name Set.member}, _) $ t $ u) =
32287
65d5c5b30747 cleaned up abstract tuple operations and named them consistently
haftmann
parents: 32135
diff changeset
   418
          infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u)
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   419
      | infer (t $ u) = infer t #> infer u
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   420
      | infer _ = I;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   421
    val new_arities = filter_out
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: 35757
diff changeset
   422
      (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 1
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   423
        | _ => false) (fold (snd #> infer) intros []);
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   424
    val params' = map (fn x =>
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   425
      (case AList.lookup op = new_arities x of
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   426
        SOME fs =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   427
          let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   428
            val T = HOLogic.dest_setT (fastype_of x);
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   429
            val Ts = HOLogic.strip_ptupleT fs T;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   430
            val x' = map_type (K (Ts ---> HOLogic.boolT)) x
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   431
          in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   432
            (x, (x',
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   433
              (HOLogic.Collect_const T $
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   434
                 HOLogic.mk_psplits fs T HOLogic.boolT x',
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   435
               list_abs (map (pair "x") Ts, HOLogic.mk_mem
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   436
                 (HOLogic.mk_ptuple fs T (map Bound (length fs downto 0)),
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   437
                  x)))))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   438
          end
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   439
       | NONE => (x, (x, (x, x))))) params;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   440
    val (params1, (params2, params3)) =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   441
      params' |> map snd |> split_list ||> split_list;
30860
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   442
    val paramTs = map fastype_of params;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   443
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   444
    (* equations for converting sets to predicates *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   445
    val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   446
      let
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   447
        val fs = the_default [] (AList.lookup op = new_arities c);
30860
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   448
        val (Us, U) = split_last (binder_types T);
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   449
        val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   450
          [Pretty.str "Argument types",
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   451
           Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)),
30860
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   452
           Pretty.str ("of " ^ s ^ " do not agree with types"),
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   453
           Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)),
30860
e5f9477aed50 Added check whether argument types of inductive set agree with types of declared
berghofe
parents: 30528
diff changeset
   454
           Pretty.str "of declared parameters"]));
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   455
        val Ts = HOLogic.strip_ptupleT fs U;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   456
        val c' = Free (s ^ "p",
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   457
          map fastype_of params1 @ Ts ---> HOLogic.boolT)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   458
      in
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   459
        ((c', (fs, U, Ts)),
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   460
         (list_comb (c, params2),
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   461
          HOLogic.Collect_const U $ HOLogic.mk_psplits fs U HOLogic.boolT
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   462
            (list_comb (c', params1))))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   463
      end) |> split_list |>> split_list;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   464
    val eqns' = eqns @
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   465
      map (prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq)
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
   466
        (mem_Collect_eq :: @{thm split_conv} :: to_pred_simps);
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   467
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   468
    (* predicate version of the introduction rules *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   469
    val intros' =
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   470
      map (fn (name_atts, t) => (name_atts,
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   471
        t |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   472
        map_aterms (fn u =>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   473
          (case AList.lookup op = params' u of
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   474
             SOME (_, (u', _)) => u'
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   475
           | NONE => u)) |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   476
        Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |>
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   477
        eta_contract (member op = cs' orf is_pred pred_arities))) intros;
30345
76fd85bbf139 more uniform handling of binding in packages;
wenzelm
parents: 30305
diff changeset
   478
    val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn;
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   479
    val monos' = map (to_pred [] (Context.Proof lthy)) monos;
38665
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 37863
diff changeset
   480
    val ({preds, intrs, elims, raw_induct, eqs, ...}, lthy1) =
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   481
      Inductive.add_ind_def
33669
ae9a2ea9a989 inductive: eliminated obsolete kind;
wenzelm
parents: 33643
diff changeset
   482
        {quiet_mode = quiet_mode, verbose = verbose, alt_name = Binding.empty,
29389
0a49f940d729 inductive: added fork_mono flag;
wenzelm
parents: 29288
diff changeset
   483
          coind = coind, no_elim = no_elim, no_ind = no_ind,
0a49f940d729 inductive: added fork_mono flag;
wenzelm
parents: 29288
diff changeset
   484
          skip_mono = skip_mono, fork_mono = fork_mono}
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   485
        cs' intros' monos' params1 cnames_syn' lthy;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   486
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   487
    (* define inductive sets using previously defined predicates *)
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   488
    val (defs, lthy2) = lthy1
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   489
      |> Local_Theory.conceal  (* FIXME ?? *)
33766
c679f05600cd adapted Local_Theory.define -- eliminated odd thm kind;
wenzelm
parents: 33671
diff changeset
   490
      |> fold_map Local_Theory.define
33278
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   491
        (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.empty_binding,
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   492
           fold_rev lambda params (HOLogic.Collect_const U $
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   493
             HOLogic.mk_psplits fs U HOLogic.boolT (list_comb (p, params3))))))
ba9f52f56356 conceal internal bindings;
wenzelm
parents: 33049
diff changeset
   494
           (cnames_syn ~~ cs_info ~~ preds))
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   495
      ||> Local_Theory.restore_naming lthy1;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   496
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   497
    (* prove theorems for converting predicate to set notation *)
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   498
    val lthy3 = fold
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   499
      (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn lthy =>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   500
        let val conv_thm =
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   501
          Goal.prove lthy (map (fst o dest_Free) params) []
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   502
            (HOLogic.mk_Trueprop (HOLogic.mk_eq
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   503
              (list_comb (p, params3),
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   504
               list_abs (map (pair "x") Ts, HOLogic.mk_mem
32342
3fabf5b5fc83 path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents: 32287
diff changeset
   505
                 (HOLogic.mk_ptuple fs U (map Bound (length fs downto 0)),
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   506
                  list_comb (c, params))))))
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   507
            (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps
37136
e0c9d3e49e15 dropped legacy theorem bindings
haftmann
parents: 36960
diff changeset
   508
              [def, mem_Collect_eq, @{thm split_conv}]) 1))
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   509
        in
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   510
          lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"),
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   511
            [Attrib.internal (K pred_set_conv_att)]),
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   512
              [conv_thm]) |> snd
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   513
        end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   514
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   515
    (* convert theorems to set notation *)
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27330
diff changeset
   516
    val rec_name =
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   517
      if Binding.is_empty alt_name then
30223
24d975352879 renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents: 30089
diff changeset
   518
        Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn))
28083
103d9282a946 explicit type Name.binding for higher-specification elements;
wenzelm
parents: 27330
diff changeset
   519
      else alt_name;
33671
4b0f2599ed48 modernized structure Local_Theory;
wenzelm
parents: 33670
diff changeset
   520
    val cnames = map (Local_Theory.full_name lthy3 o #1) cnames_syn;  (* FIXME *)
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   521
    val (intr_names, intr_atts) = split_list (map fst intros);
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   522
    val raw_induct' = to_set [] (Context.Proof lthy3) raw_induct;
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37677
diff changeset
   523
    val (intrs', elims', eqs', induct, inducts, lthy4) =
35757
c2884bec5463 adding Spec_Rules to definitional package inductive and inductive_set
bulwahn
parents: 35646
diff changeset
   524
      Inductive.declare_rules rec_name coind no_ind cnames (map fst defs)
33459
wenzelm
parents: 33458
diff changeset
   525
        (map (to_set [] (Context.Proof lthy3)) intrs) intr_names intr_atts
wenzelm
parents: 33458
diff changeset
   526
        (map (fn th => (to_set [] (Context.Proof lthy3) th,
34986
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34903
diff changeset
   527
           map fst (fst (Rule_Cases.get th)),
7f7939c9370f Added "constraints" tag / attribute for specifying the number of equality
berghofe
parents: 34903
diff changeset
   528
           Rule_Cases.get_constraints th)) elims)
38665
e92223c886f8 introducing simplification equations for inductive sets; added data structure for storing equations; rewriting retrieval of simplification equation for inductive predicates and sets
bulwahn
parents: 37863
diff changeset
   529
        (map (to_set [] (Context.Proof lthy3)) eqs) raw_induct' lthy3;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   530
  in
35646
b32d6c1bdb4d Added inducts field to inductive_result.
berghofe
parents: 35364
diff changeset
   531
    ({intrs = intrs', elims = elims', induct = induct, inducts = inducts,
37734
489ac1ecb9f1 added the new command inductive_cases to derive simplification equations for inductive predicates; added binding simps for general simplification equation
bulwahn
parents: 37677
diff changeset
   532
      raw_induct = raw_induct', preds = map fst defs, eqs = eqs'},
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   533
     lthy4)
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   534
  end;
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   535
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   536
val add_inductive_i = Inductive.gen_add_inductive_i add_ind_set_def;
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   537
val add_inductive = Inductive.gen_add_inductive add_ind_set_def;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   538
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   539
val mono_add_att = to_pred_att [] #> Inductive.mono_add;
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   540
val mono_del_att = to_pred_att [] #> Inductive.mono_del;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   541
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   542
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   543
(** package setup **)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   544
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   545
(* setup theory *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   546
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   547
val setup =
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   548
  Attrib.setup @{binding pred_set_conv} (Scan.succeed pred_set_conv_att)
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   549
    "declare rules for converting between predicate and set notation" #>
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   550
  Attrib.setup @{binding to_set} (Attrib.thms >> to_set_att)
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   551
    "convert rule to set notation" #>
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   552
  Attrib.setup @{binding to_pred} (Attrib.thms >> to_pred_att)
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   553
    "convert rule to predicate notation" #>
31998
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31723
diff changeset
   554
  Attrib.setup @{binding code_ind_set}
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31723
diff changeset
   555
    (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) >> code_ind_att))
2c7a24f74db9 code attributes use common underscore convention
haftmann
parents: 31723
diff changeset
   556
    "introduction rules for executable predicates" #>
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   557
  Codegen.add_preprocessor codegen_preproc #>
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   558
  Attrib.setup @{binding mono_set} (Attrib.add_del mono_add_att mono_del_att)
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   559
    "declaration of monotonicity rule for set operators" #>
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   560
  Context.theory_map (Simplifier.map_ss (fn ss => ss addsimprocs [collect_mem_simproc]));
7173bf123335 simplified attribute setup;
wenzelm
parents: 30345
diff changeset
   561
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   562
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   563
(* outer syntax *)
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   564
31723
f5cafe803b55 discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents: 30860
diff changeset
   565
val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def;
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   566
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24815
diff changeset
   567
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
   568
  Outer_Syntax.local_theory' "inductive_set" "define inductive sets" Keyword.thy_decl
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   569
    (ind_set_decl false);
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   570
24867
e5b55d7be9bb simplified interfaces for outer syntax;
wenzelm
parents: 24815
diff changeset
   571
val _ =
36960
01594f816e3a prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents: 36945
diff changeset
   572
  Outer_Syntax.local_theory' "coinductive_set" "define coinductive sets" Keyword.thy_decl
33458
ae1f5d89b082 proper naming convention lthy: local_theory, but ctxt: Proof.context for arbitrary context;
wenzelm
parents: 33368
diff changeset
   573
    (ind_set_decl true);
23764
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   574
15f81c5d5330 New wrapper for defining inductive sets with new inductive
berghofe
parents:
diff changeset
   575
end;