src/ZF/ind_syntax.ML
author immler@in.tum.de
Thu, 26 Feb 2009 10:13:43 +0100
changeset 30151 629f3a92863e
parent 28965 1de908189869
child 30345 76fd85bbf139
permissions -rw-r--r--
removed global ref dfg_format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12243
wenzelm
parents: 8819
diff changeset
     1
(*  Title:      ZF/ind_syntax.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
12243
wenzelm
parents: 8819
diff changeset
     6
Abstract Syntax functions for Inductive Definitions.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
     9
structure Ind_Syntax =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    10
struct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
4804
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    12
(*Print tracing messages during processing of "inductive" theory sections*)
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    13
val trace = ref false;
02b7c759159b Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents: 4352
diff changeset
    14
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    15
fun traceIt msg thy t =
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26189
diff changeset
    16
  if !trace then (tracing (msg ^ Syntax.string_of_term_global thy t); t)
17988
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    17
  else t;
47f81afce1b4 traceIt: plain term;
wenzelm
parents: 16867
diff changeset
    18
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    19
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    20
(** Abstract syntax definitions for ZF **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    22
val iT = Type("i",[]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    25
fun mk_all_imp (A,P) =
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    26
    FOLogic.all_const iT $
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    27
      Abs("v", iT, FOLogic.imp $ (@{const mem} $ Bound 0 $ A) $
18176
ae9bd644d106 Term.betapply;
wenzelm
parents: 17988
diff changeset
    28
	           Term.betapply(P, Bound 0));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    30
fun mk_Collect (a, D, t) = @{const Collect} $ D $ absfree (a, iT, t);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    32
(*simple error-checking in the premises of an inductive definition*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    33
fun chk_prem rec_hd (Const (@{const_name "op &"}, _) $ _ $ _) =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    34
        error"Premises may not be conjuctive"
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    35
  | chk_prem rec_hd (Const (@{const_name mem}, _) $ t $ X) =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 21539
diff changeset
    36
        (Logic.occs(rec_hd,t) andalso error "Recursion term on left of member symbol"; ())
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    37
  | chk_prem rec_hd t =
22567
1565d476a9e2 removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents: 21539
diff changeset
    38
        (Logic.occs(rec_hd,t) andalso error "Recursion term in side formula"; ());
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    39
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    40
(*Return the conclusion of a rule, of the form t:X*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    41
fun rule_concl rl =
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    42
    let val Const (@{const_name Trueprop}, _) $ (Const (@{const_name mem}, _) $ t $ X) =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    43
                Logic.strip_imp_concl rl
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    44
    in  (t,X)  end;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    45
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    46
(*As above, but return error message if bad*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    47
fun rule_concl_msg sign rl = rule_concl rl
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    48
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26189
diff changeset
    49
                          Syntax.string_of_term_global sign rl);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
(*For deriving cases rules.  CollectD2 discards the domain, which is redundant;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    52
  read_instantiate replaces a propositional variable by a formula variable*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    53
val equals_CollectD =
27239
f2f42f9fa09d pervasive RuleInsts;
wenzelm
parents: 27230
diff changeset
    54
    read_instantiate @{context} [(("W", 0), "?Q")]
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24826
diff changeset
    55
        (make_elim (@{thm equalityD1} RS @{thm subsetD} RS @{thm CollectD2}));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    56
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    58
(** For datatype definitions **)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    59
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    60
(*Constructor name, type, mixfix info;
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    61
  internal name from mixfix, datatype sets, full premises*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    62
type constructor_spec =
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    63
    (string * typ * mixfix) * string * term list * term list;
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 4972
diff changeset
    64
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    65
fun dest_mem (Const (@{const_name mem}, _) $ x $ A) = (x, A)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    66
  | dest_mem _ = error "Constructor specifications must have the form x:A";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    67
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    68
(*read a constructor specification*)
27261
5b3101338f42 eliminated old Sign.read_term/Thm.read_cterm etc.;
wenzelm
parents: 27239
diff changeset
    69
fun read_construct ctxt (id, sprems, syn) =
5b3101338f42 eliminated old Sign.read_term/Thm.read_cterm etc.;
wenzelm
parents: 27239
diff changeset
    70
    let val prems = map (Syntax.parse_term ctxt #> TypeInfer.constrain FOLogic.oT) sprems
5b3101338f42 eliminated old Sign.read_term/Thm.read_cterm etc.;
wenzelm
parents: 27239
diff changeset
    71
          |> Syntax.check_terms ctxt
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    72
        val args = map (#1 o dest_mem) prems
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    73
        val T = (map (#2 o dest_Free) args) ---> iT
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    74
                handle TERM _ => error
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    75
                    "Bad variable in constructor specification"
568
756b0e2a6cac replaced Lexicon.scan_id by Scanner.scan_id;
wenzelm
parents: 543
diff changeset
    76
        val name = Syntax.const_name id syn  (*handle infix constructors*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    77
    in ((id,T,syn), name, args, prems) end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    78
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    79
val read_constructs = map o map o read_construct;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    81
(*convert constructor specifications into introduction rules*)
3925
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    82
fun mk_intr_tms sg (rec_tm, constructs) =
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    83
  let
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    84
    fun mk_intr ((id,T,syn), name, args, prems) =
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    85
      Logic.list_implies
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    86
        (map FOLogic.mk_Trueprop prems,
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    87
	 FOLogic.mk_Trueprop
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 27261
diff changeset
    88
	    (@{const mem} $ list_comb (Const (Sign.full_bname sg name, T), args)
4352
7ac9f3e8a97d Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents: 3925
diff changeset
    89
	               $ rec_tm))
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    90
  in  map mk_intr constructs  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    91
3925
90f499226ab9 (co) inductive / datatype package adapted to qualified names;
wenzelm
parents: 2266
diff changeset
    92
fun mk_all_intr_tms sg arg = List.concat (ListPair.map (mk_intr_tms sg) arg);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
    94
fun mk_Un (t1, t2) = @{const Un} $ t1 $ t2;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    96
(*Make a datatype's domain: form the union of its set parameters*)
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    97
fun union_params (rec_tm, cs) =
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    98
  let val (_,args) = strip_comb rec_tm
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    99
      fun is_ind arg = (type_of arg = iT)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 13150
diff changeset
   100
  in  case List.filter is_ind (args @ cs) of
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   101
         []     => @{const 0}
23419
8c30dd4b3b22 BalancedTree;
wenzelm
parents: 23156
diff changeset
   102
       | u_args => BalancedTree.make mk_Un u_args
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   103
  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   104
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   106
(*Includes rules for succ and Pair since they are common constructions*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   107
val elim_rls =
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   108
  [@{thm asm_rl}, @{thm FalseE}, @{thm succ_neq_0}, @{thm sym} RS @{thm succ_neq_0},
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   109
   @{thm Pair_neq_0}, @{thm sym} RS @{thm Pair_neq_0}, @{thm Pair_inject},
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   110
   make_elim @{thm succ_inject}, @{thm refl_thin}, @{thm conjE}, @{thm exE}, @{thm disjE}];
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   111
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   112
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   113
(*From HOL/ex/meson.ML: raises exception if no rules apply -- unlike RL*)
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   114
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls))
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   115
  | tryres (th, []) = raise THM("tryres", 0, [th]);
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   116
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   117
fun gen_make_elim elim_rls rl =
7694
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   118
      standard (tryres (rl, elim_rls @ [revcut_rl]));
20121c9dc1a6 tryres, gen_make_elim moved here;
wenzelm
parents: 6112
diff changeset
   119
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   120
(*Turns iff rules into safe elimination rules*)
26189
9808cca5c54d misc cleanup of embedded ML code;
wenzelm
parents: 26059
diff changeset
   121
fun mk_free_SEs iffs = map (gen_make_elim [@{thm conjE}, @{thm FalseE}]) (iffs RL [@{thm iffD1}]);
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   122
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   123
end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   124