src/ZF/ind_syntax.ML
author wenzelm
Wed, 15 Oct 1997 15:12:59 +0200
changeset 3872 a5839ecee7b8
parent 2266 82aef6857c5b
child 3925 90f499226ab9
permissions -rw-r--r--
tuned; prepare ext;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
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
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Abstract Syntax functions for Inductive Definitions
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
(*The structure protects these items from redeclaration (somewhat!).  The 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    10
  datatype definitions in theory files refer to these items by name!
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    11
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    12
structure Ind_Syntax =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    13
struct
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
(** Abstract syntax definitions for FOL and ZF **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
val iT = Type("i",[])
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
and oT = Type("o",[]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
1738
a70a5bc5e315 moved ap_split to cartprod.ML
paulson
parents: 1461
diff changeset
    20
a70a5bc5e315 moved ap_split to cartprod.ML
paulson
parents: 1461
diff changeset
    21
(** Logical constants **)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
val conj = Const("op &", [oT,oT]--->oT)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
and disj = Const("op |", [oT,oT]--->oT)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
and imp = Const("op -->", [oT,oT]--->oT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
val eq_const = Const("op =", [iT,iT]--->oT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
val mem_const = Const("op :", [iT,iT]--->oT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
val exists_const = Const("Ex", [iT-->oT]--->oT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
fun mk_exists (Free(x,T),P) = exists_const $ (absfree (x,T,P));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
val all_const = Const("All", [iT-->oT]--->oT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
fun mk_all (Free(x,T),P) = all_const $ (absfree (x,T,P));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    36
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
(*Creates All(%v.v:A --> P(v)) rather than Ball(A,P) *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
fun mk_all_imp (A,P) = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
    all_const $ Abs("v", iT, imp $ (mem_const $ Bound 0 $ A) $ (P $ Bound 0));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
val Part_const = Const("Part", [iT,iT-->iT]--->iT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
val Collect_const = Const("Collect", [iT,iT-->oT]--->iT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
fun mk_Collect (a,D,t) = Collect_const $ D $ absfree(a, iT, t);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    46
val Trueprop = Const("Trueprop",oT-->propT);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    47
fun mk_tprop P = Trueprop $ P;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    49
(*simple error-checking in the premises of an inductive definition*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    50
fun chk_prem rec_hd (Const("op &",_) $ _ $ _) =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    51
        error"Premises may not be conjuctive"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    52
  | chk_prem rec_hd (Const("op :",_) $ t $ X) = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    53
        deny (Logic.occs(rec_hd,t)) "Recursion term on left of member symbol"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    54
  | chk_prem rec_hd t = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    55
        deny (Logic.occs(rec_hd,t)) "Recursion term in side formula";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    56
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    57
(*Return the conclusion of a rule, of the form t:X*)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
fun rule_concl rl = 
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    59
    let val Const("Trueprop",_) $ (Const("op :",_) $ t $ X) = 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    60
                Logic.strip_imp_concl rl
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    61
    in  (t,X)  end;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    62
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    63
(*As above, but return error message if bad*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    64
fun rule_concl_msg sign rl = rule_concl rl
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 231
diff changeset
    65
    handle Bind => error ("Ill-formed conclusion of introduction rule: " ^ 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    66
                          Sign.string_of_term sign rl);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    67
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
(*For deriving cases rules.  CollectD2 discards the domain, which is redundant;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    69
  read_instantiate replaces a propositional variable by a formula variable*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    70
val equals_CollectD = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
    read_instantiate [("W","?Q")]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
        (make_elim (equalityD1 RS subsetD RS CollectD2));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    75
(** For datatype definitions **)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    76
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    77
fun dest_mem (Const("op :",_) $ x $ A) = (x,A)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    78
  | dest_mem _ = error "Constructor specifications must have the form x:A";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    79
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    80
(*read a constructor specification*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    81
fun read_construct sign (id, sprems, syn) =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    82
    let val prems = map (readtm sign oT) sprems
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    83
        val args = map (#1 o dest_mem) prems
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    84
        val T = (map (#2 o dest_Free) args) ---> iT
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    85
                handle TERM _ => error 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    86
                    "Bad variable in constructor specification"
568
756b0e2a6cac replaced Lexicon.scan_id by Scanner.scan_id;
wenzelm
parents: 543
diff changeset
    87
        val name = Syntax.const_name id syn  (*handle infix constructors*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    88
    in ((id,T,syn), name, args, prems) end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    89
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    90
val read_constructs = map o map o read_construct;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    92
(*convert constructor specifications into introduction rules*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    93
fun mk_intr_tms (rec_tm, constructs) =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    94
  let fun mk_intr ((id,T,syn), name, args, prems) =
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    95
          Logic.list_implies
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    96
              (map mk_tprop prems,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
    97
               mk_tprop (mem_const $ list_comb(Const(name,T), args) $ rec_tm)) 
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    98
  in  map mk_intr constructs  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
    99
2266
82aef6857c5b Replaced map...~~ by ListPair.map
paulson
parents: 2226
diff changeset
   100
fun mk_all_intr_tms arg = List.concat (ListPair.map mk_intr_tms arg);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   102
val Un          = Const("op Un", [iT,iT]--->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   103
and empty       = Const("0", iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   104
and univ        = Const("univ", iT-->iT)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   105
and quniv       = Const("quniv", iT-->iT);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   107
(*Make a datatype's domain: form the union of its set parameters*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   108
fun union_params rec_tm =
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   109
  let val (_,args) = strip_comb rec_tm
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   110
  in  case (filter (fn arg => type_of arg = iT) args) of
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   111
         []    => empty
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   112
       | iargs => fold_bal (app Un) iargs
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   113
  end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   114
742
faa3efc1d130 data_domain,Codata_domain: removed replicate; now return one
lcp
parents: 578
diff changeset
   115
(*Previously these both did    replicate (length rec_tms);  however now
faa3efc1d130 data_domain,Codata_domain: removed replicate; now return one
lcp
parents: 578
diff changeset
   116
  [q]univ itself constitutes the sum domain for mutual recursion!*)
faa3efc1d130 data_domain,Codata_domain: removed replicate; now return one
lcp
parents: 578
diff changeset
   117
fun data_domain rec_tms = univ $ union_params (hd rec_tms);
faa3efc1d130 data_domain,Codata_domain: removed replicate; now return one
lcp
parents: 578
diff changeset
   118
fun Codata_domain rec_tms = quniv $ union_params (hd rec_tms);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   119
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
(*Could go to FOL, but it's hardly general*)
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   121
val def_swap_iff = prove_goal IFOL.thy "a==b ==> a=c <-> c=b"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   122
 (fn [def] => [(rewtac def), (rtac iffI 1), (REPEAT (etac sym 1))]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   123
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
val def_trans = prove_goal IFOL.thy "[| f==g;  g(a)=b |] ==> f(a)=b"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
  (fn [rew,prem] => [ rewtac rew, rtac prem 1 ]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   126
55
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   127
(*Delete needless equality assumptions*)
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   128
val refl_thin = prove_goal IFOL.thy "!!P. [| a=a;  P |] ==> P"
331d93292ee0 ZF/ind-syntax/refl_thin: new
lcp
parents: 14
diff changeset
   129
     (fn _ => [assume_tac 1]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   131
(*Includes rules for succ and Pair since they are common constructions*)
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   132
val elim_rls = [asm_rl, FalseE, succ_neq_0, sym RS succ_neq_0, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   133
                Pair_neq_0, sym RS Pair_neq_0, Pair_inject,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   134
                make_elim succ_inject, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1418
diff changeset
   135
                refl_thin, conjE, exE, disjE];
1418
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   136
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   137
(*Turns iff rules into safe elimination rules*)
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   138
fun mk_free_SEs iffs = map (gen_make_elim [conjE,FalseE]) (iffs RL [iffD1]);
f5f97ee67cbb Improving space efficiency of inductive/datatype definitions.
paulson
parents: 742
diff changeset
   139
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   140
end;
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 466
diff changeset
   141