Quot/quotient_term.ML
author Christian Urban <urbanc@in.tum.de>
Sat, 26 Dec 2009 08:06:45 +0100
changeset 791 fb4bfbb1a291
parent 790 3a48ffcf0f9a
child 792 a31cf260eeb3
permissions -rw-r--r--
commeted the absrep function
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
signature QUOTIENT_TERM =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
sig
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
     3
   exception LIFT_MATCH of string
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
     4
 
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
     5
   datatype flag = absF | repF
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
     6
   
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
     7
   val absrep_fun: flag -> Proof.context -> (typ * typ) -> term
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
     8
   val absrep_fun_chk: flag -> Proof.context -> (typ * typ) -> term
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
     9
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
    10
   val regularize_trm: Proof.context -> (term * term) -> term
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
    11
   val regularize_trm_chk: Proof.context -> (term * term) -> term
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
    12
   
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    13
   val inj_repabs_trm: Proof.context -> (term * term) -> term
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
    14
   val inj_repabs_trm_chk: Proof.context -> (term * term) -> term
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
end;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
structure Quotient_Term: QUOTIENT_TERM =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
struct
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    20
open Quotient_Info;
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    21
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    22
exception LIFT_MATCH of string
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    23
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    24
(******************************)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    25
(* Aggregate Rep/Abs Function *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    26
(******************************)
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
    27
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    28
(* The flag repF is for types in negative position; absF is for types *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    29
(* in positive position. Because of this, function types need to be   *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    30
(* treated specially, since there the polarity changes.               *)
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    31
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    32
datatype flag = absF | repF
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    33
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    34
fun negF absF = repF
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    35
  | negF repF = absF
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    36
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    37
val mk_id = Const (@{const_name "id"}, dummyT)
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    38
fun mk_identity ty = Const (@{const_name "id"}, ty --> ty)
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    39
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    40
(* makes a Free out of a TVar *)
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    41
fun mk_Free (TVar ((x, i), _)) = Free (unprefix "'" x ^ string_of_int i, dummyT)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    42
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    43
fun mk_compose flag (trm1, trm2) = 
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    44
  case flag of
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    45
    absF => Const (@{const_name "comp"}, dummyT) $ trm1 $ trm2
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    46
  | repF => Const (@{const_name "comp"}, dummyT) $ trm2 $ trm1
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    47
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    48
fun get_mapfun ctxt s =
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    49
let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
    50
  val thy = ProofContext.theory_of ctxt
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    51
  val exc = LIFT_MATCH ("No map function for type " ^ (quote s) ^ " found.")
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    52
  val mapfun = #mapfun (maps_lookup thy s) handle NotFound => raise exc
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    53
in
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
    54
  Const (mapfun, dummyT)
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    55
end
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    56
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    57
(* produces an aggregate map function for the       *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    58
(* rty-part of a quotient definition; abstracts     *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    59
(* over all variables listed in vs (these variables *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    60
(* correspond to the type variables in rty          *)        
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    61
fun mk_mapfun ctxt vs ty =
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    62
let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    63
  val vs' = map (mk_Free) vs
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    64
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    65
  fun mk_mapfun_aux ty =
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    66
    case ty of
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    67
      TVar _ => mk_Free ty
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    68
    | Type (_, []) => mk_id
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    69
    | Type (s, tys) => list_comb (get_mapfun ctxt s, map mk_mapfun_aux tys)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    70
    | _ => raise LIFT_MATCH ("mk_mapfun_aux (default)")
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    71
in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    72
  fold_rev Term.lambda vs' (mk_mapfun_aux ty)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    73
end
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    74
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    75
(* looks up the (varified) rty and qty for *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    76
(* a quotient definition                   *)
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    77
fun get_rty_qty ctxt s =
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
    78
let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
    79
  val thy = ProofContext.theory_of ctxt
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    80
  val exc = LIFT_MATCH ("No quotient type " ^ (quote s) ^ " found.")
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    81
  val qdata = (quotdata_lookup thy s) handle NotFound => raise exc
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    82
in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    83
  (#rtyp qdata, #qtyp qdata)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    84
end
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    85
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    86
(* receives two type-environments and looks *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    87
(* up in both of them the variable v        *)
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    88
fun double_lookup rtyenv qtyenv v =
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    89
let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    90
  val v' = fst (dest_TVar v)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    91
in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    92
  (snd (the (Vartab.lookup rtyenv v')), snd (the (Vartab.lookup qtyenv v')))
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    93
end
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    94
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
    95
(* produces the rep or abs constants for a qty *)
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    96
fun absrep_const flag ctxt qty_str =
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    97
let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    98
  val thy = ProofContext.theory_of ctxt
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
    99
  val qty_name = Long_Name.base_name qty_str
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   100
in
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   101
  case flag of
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   102
    absF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT)
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   103
  | repF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT)
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   104
end
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   105
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   106
(* produces the aggregate absrep function                          *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   107
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   108
(* - In case of function types and TFrees, we recurse, taking in   *) 
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   109
(*   the first case the polarity change into account.              *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   110
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   111
(* - If the type constructors are equal, we recurse for the        *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   112
(*   arguments and prefix the appropriate map function             *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   113
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   114
(* - If the type constructors are unequal, there must be an        *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   115
(*   instance of quotient types:                                   *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   116
(*     - we first look up the corresponding rty_pat and qty_pat    *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   117
(*       from the quotient definition; the arguments of qty_pat    *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   118
(*       must be some distinct TVars                               *)  
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   119
(*     - we then match the rty_pat with rty and qty_pat with qty;  *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   120
(*       if matching fails the types do not match                  *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   121
(*     - the matching produces two environments, we look up the    *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   122
(*       assignments for the qty_pat variables and recurse on the  *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   123
(*       assignmetnts                                              *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   124
(*     - we prefix the aggregate map function for the rty_pat,     *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   125
(*       which is an abstraction over all type variables           *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   126
(*     - finally we compse the result with the appropriate         *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   127
(*       absrep function                                           *)    
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   128
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   129
(*   The composition is necessary for types like                   *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   130
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   131
(*      ('a list) list / ('a foo) foo                              *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   132
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   133
(*   The matching is necessary for types like                      *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   134
(*                                                                 *)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   135
(*      ('a * 'a) list / 'a bar                                    *)  
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   136
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   137
fun absrep_fun flag ctxt (rty, qty) =
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   138
  if rty = qty  
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   139
  then mk_identity qty
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   140
  else
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   141
    case (rty, qty) of
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   142
      (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   143
        let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   144
          val arg1 = absrep_fun (negF flag) ctxt (ty1, ty1')
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   145
          val arg2 = absrep_fun flag ctxt (ty2, ty2')
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   146
        in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   147
          list_comb (get_mapfun ctxt "fun", [arg1, arg2])
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   148
        end
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   149
    | (Type (s, tys), Type (s', tys')) =>
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   150
        if s = s'
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   151
        then 
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   152
           let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   153
             val args = map (absrep_fun flag ctxt) (tys ~~ tys')
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   154
           in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   155
             list_comb (get_mapfun ctxt s, args)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   156
           end
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   157
        else
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   158
           let
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   159
             val thy = ProofContext.theory_of ctxt
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   160
             val exc = (LIFT_MATCH "absrep_fun (Types do not match.)")
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   161
             val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s'
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   162
             val (rtyenv, qtyenv) = 
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   163
                     (Sign.typ_match thy (rty_pat, rty) Vartab.empty,
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   164
                      Sign.typ_match thy (qty_pat, qty) Vartab.empty)
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   165
                      handle MATCH_TYPE => raise exc
790
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   166
             val args_aux = map (double_lookup rtyenv qtyenv) vs            
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   167
             val args = map (absrep_fun flag ctxt) args_aux
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   168
             val map_fun = mk_mapfun ctxt vs rty_pat       
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   169
             val result = list_comb (map_fun, args) 
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   170
           in
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   171
             mk_compose flag (absrep_const flag ctxt s', result)
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   172
           end 
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   173
    | (TFree x, TFree x') =>
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   174
        if x = x'
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   175
        then mk_identity qty
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   176
        else raise (LIFT_MATCH "absrep_fun (frees)")
3a48ffcf0f9a generalised absrep function; needs consolidation
Christian Urban <urbanc@in.tum.de>
parents: 786
diff changeset
   177
    | (TVar _, TVar _) => raise (LIFT_MATCH "absrep_fun (vars)")
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   178
    | _ => raise (LIFT_MATCH "absrep_fun (default)")
774
b4ffb8826105 moved get_fun into quotient_term; this simplifies the overall including structure of the package
Christian Urban <urbanc@in.tum.de>
parents: 762
diff changeset
   179
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   180
fun absrep_fun_chk flag ctxt (rty, qty) =
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   181
  absrep_fun flag ctxt (rty, qty)
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   182
  |> Syntax.check_term ctxt
791
fb4bfbb1a291 commeted the absrep function
Christian Urban <urbanc@in.tum.de>
parents: 790
diff changeset
   183
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   184
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   185
(* Regularizing an rtrm means:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
 
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   187
 - Quantifiers over types that need lifting are replaced 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
   by bounded quantifiers, for example:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   190
      All P  ----> All (Respects R) P
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   192
   where the aggregate relation R is given by the rty and qty;
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
 
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   194
 - Abstractions over types that need lifting are replaced
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   195
   by bounded abstractions, for example:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
      
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   197
      %x. P  ----> Ball (Respects R) %x. P
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   199
 - Equalities over types that need lifting are replaced by
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   200
   corresponding equivalence relations, for example:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   202
      A = B  ----> R A B
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
   or 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   206
      A = B  ----> (R ===> R) A B
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
   for more complicated types of A and B
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
*)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
783
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   211
(* instantiates TVars so that the term is of type ty *)
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   212
fun force_typ thy trm ty =
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   213
let
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   214
  val trm_ty = fastype_of trm
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   215
  val ty_inst = Sign.typ_match thy (trm_ty, ty) Vartab.empty
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   216
in
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   217
  map_types (Envir.subst_type ty_inst) trm
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   218
end
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   219
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   220
(* builds the aggregate equivalence relation *)
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   221
(* that will be the argument of Respects     *)
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   222
fun mk_resp_arg ctxt (rty, qty) =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   224
  val thy = ProofContext.theory_of ctxt
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
in  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  if rty = qty
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  then HOLogic.eq_const rty
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  else
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
    case (rty, qty) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
      (Type (s, tys), Type (s', tys')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
       if s = s' 
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   232
       then 
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   233
         let
785
Christian Urban <urbanc@in.tum.de>
parents: 784
diff changeset
   234
           val exc = LIFT_MATCH ("mk_resp_arg (no relation map function found for type " ^ s ^ ")") 
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   235
           val relmap = #relmap (maps_lookup thy s) handle NotFound => raise exc
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   236
           val args = map (mk_resp_arg ctxt) (tys ~~ tys')
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   237
         in
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   238
           list_comb (Const (relmap, dummyT), args) 
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   239
         end  
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   240
       else 
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   241
         let
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   242
           val exc = LIFT_MATCH ("mk_resp_arg (no quotient found for type " ^ s ^ ")") 
786
d6407afb913c used Local_Theory.declaration for storing quotdata
Christian Urban <urbanc@in.tum.de>
parents: 785
diff changeset
   243
           val equiv_rel = #equiv_rel (quotdata_lookup thy s') handle NotFound => raise exc
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   244
           (* FIXME: check in this case that the rty and qty *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   245
           (* FIXME: correspond to each other *)
783
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   246
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   247
           (* we need to instantiate the TVars in the relation *)
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   248
           val thy = ProofContext.theory_of ctxt 
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   249
           val forced_equiv_rel = force_typ thy equiv_rel (rty --> rty --> @{typ bool})
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   250
         in
784
da75568e7f12 renamed some fields in the info records
Christian Urban <urbanc@in.tum.de>
parents: 783
diff changeset
   251
           forced_equiv_rel
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   252
         end
783
06e17083e90b modified mk_resp_arg so that the user can give terms as equivalence relations, not just constants
Christian Urban <urbanc@in.tum.de>
parents: 782
diff changeset
   253
      | _ => HOLogic.eq_const rty
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
             (* FIXME: check that the types correspond to each other? *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
val mk_babs = Const (@{const_name Babs}, dummyT)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
val mk_ball = Const (@{const_name Ball}, dummyT)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
val mk_bex  = Const (@{const_name Bex}, dummyT)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
val mk_resp = Const (@{const_name Respects}, dummyT)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
(* - applies f to the subterm of an abstraction,   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
(*   otherwise to the given term,                  *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
(* - used by regularize, therefore abstracted      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
(*   variables do not have to be treated specially *)
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   266
fun apply_subt f (trm1, trm2) =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  case (trm1, trm2) of
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   268
    (Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t'))
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   269
  | _ => f (trm1, trm2)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
(* the major type of All and Ex quantifiers *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
fun qnt_typ ty = domain_type (domain_type ty)  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
(* produces a regularized version of rtrm       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
(*                                              *)
785
Christian Urban <urbanc@in.tum.de>
parents: 784
diff changeset
   277
(* - the result might contain dummyTs           *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
(*                                              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
(* - for regularisation we do not need any      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
(*   special treatment of bound variables       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   282
fun regularize_trm ctxt (rtrm, qtrm) =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  case (rtrm, qtrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
    (Abs (x, ty, t), Abs (_, ty', t')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
       let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   286
         val subtrm = Abs(x, ty, regularize_trm ctxt (t, t'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
         if ty = ty' then subtrm
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   289
         else mk_babs $ (mk_resp $ mk_resp_arg ctxt (ty, ty')) $ subtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  | (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
       let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   294
         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
         if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   297
         else mk_ball $ (mk_resp $ mk_resp_arg ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  | (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
       let
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   302
         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
         if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   305
         else mk_bex $ (mk_resp $ mk_resp_arg ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  | (* equalities need to be replaced by appropriate equivalence relations *) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
    (Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
         if ty = ty' then rtrm
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   311
         else mk_resp_arg ctxt (domain_type ty, domain_type ty') 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   313
  | (* in this case we just check whether the given equivalence relation is correct *) 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
    (rel, Const (@{const_name "op ="}, ty')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
       let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
         val exc = LIFT_MATCH "regularise (relation mismatch)"
785
Christian Urban <urbanc@in.tum.de>
parents: 784
diff changeset
   317
         val rel_ty = fastype_of rel
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   318
         val rel' = mk_resp_arg ctxt (domain_type rel_ty, domain_type ty') 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
       in 
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   320
         if rel' aconv rel then rtrm else raise exc
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
       end  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  | (_, Const _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
       let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
         fun same_name (Const (s, T)) (Const (s', T')) = (s = s') (*andalso (T = T')*)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
           | same_name _ _ = false
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
          (* TODO/FIXME: This test is not enough. *) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
          (*             Why?                     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
          (* Because constants can have the same name but not be the same
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
             constant.  All overloaded constants have the same name but because
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
             of different types they do differ.
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
        
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
             This code will let one write a theorem where plus on nat is
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
             matched to plus on int, even if the latter is defined differently.
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
             This would result in hard to understand failures in injection and
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
             cleaning. *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
           (* cu: if I also test the type, then something else breaks *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
         if same_name rtrm qtrm then rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
         else 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
           let 
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   343
             val thy = ProofContext.theory_of ctxt
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   344
             val qtrm_str = Syntax.string_of_term ctxt qtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
             val exc1 = LIFT_MATCH ("regularize (constant " ^ qtrm_str ^ " not found)")
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
             val exc2 = LIFT_MATCH ("regularize (constant " ^ qtrm_str ^ " mismatch)")
785
Christian Urban <urbanc@in.tum.de>
parents: 784
diff changeset
   347
             val rtrm' = #rconst (qconsts_lookup thy qtrm) handle NotFound => raise exc1
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
           in 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
             if Pattern.matches thy (rtrm', rtrm) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
             then rtrm else raise exc2
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
           end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
       end 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
  | (t1 $ t2, t1' $ t2') =>
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   355
       (regularize_trm ctxt (t1, t1')) $ (regularize_trm ctxt (t2, t2'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
  | (Bound i, Bound i') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
       if i = i' then rtrm 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
       else raise (LIFT_MATCH "regularize (bounds mismatch)")
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
  | _ =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
       let 
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   363
         val rtrm_str = Syntax.string_of_term ctxt rtrm
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   364
         val qtrm_str = Syntax.string_of_term ctxt qtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
         raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")"))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   369
fun regularize_trm_chk ctxt (rtrm, qtrm) =
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   370
  regularize_trm ctxt (rtrm, qtrm) 
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   371
  |> Syntax.check_term ctxt
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
(*
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   374
Injection of Rep/Abs means:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
  For abstractions
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
:
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   378
  * If the type of the abstraction needs lifting, then we add Rep/Abs 
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   379
    around the abstraction; otherwise we leave it unchanged.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  For applications:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
  
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   383
  * If the application involves a bounded quantifier, we recurse on 
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   384
    the second argument. If the application is a bounded abstraction,
782
86c7ed9f354f cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents: 779
diff changeset
   385
    we always put an Rep/Abs around it (since bounded abstractions
86c7ed9f354f cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents: 779
diff changeset
   386
    are assumed to always need lifting). Otherwise we recurse on both 
86c7ed9f354f cleaed a bit function mk_typedef_main
Christian Urban <urbanc@in.tum.de>
parents: 779
diff changeset
   387
    arguments.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   389
  For constants:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   391
  * If the constant is (op =), we leave it always unchanged. 
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   392
    Otherwise the type of the constant needs lifting, we put
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   393
    and Rep/Abs around it. 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   395
  For free variables:
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   396
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   397
  * We put aRep/Abs around it if the type needs lifting.
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   398
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   399
  Vars case cannot occur.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
*)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   402
fun mk_repabs ctxt (T, T') trm = 
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   403
  absrep_fun repF ctxt (T, T') $ (absrep_fun absF ctxt (T, T') $ trm)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
(* bound variables need to be treated properly,     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
(* as the type of subterms needs to be calculated   *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   409
fun inj_repabs_trm ctxt (rtrm, qtrm) =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
 case (rtrm, qtrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   412
       Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   415
       Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  | (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
        val qty = fastype_of qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
      in
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   422
        mk_repabs ctxt (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm ctxt (t, t')))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
      end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
  | (Abs (x, T, t), Abs (x', T', t')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
        val qty = fastype_of qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
        val (y, s) = Term.dest_abs (x, T, t)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
        val (_, s') = Term.dest_abs (x', T', t')
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
        val yvar = Free (y, T)
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   432
        val result = Term.lambda_name (y, yvar) (inj_repabs_trm ctxt (s, s'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
      in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
        if rty = qty then result
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   435
        else mk_repabs ctxt (rty, qty) result
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
      end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
  | (t $ s, t' $ s') =>  
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   439
       (inj_repabs_trm ctxt (t, t')) $ (inj_repabs_trm ctxt (s, s'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
  | (Free (_, T), Free (_, T')) => 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
        if T = T' then rtrm 
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   443
        else mk_repabs ctxt (T, T') rtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  | (_, Const (@{const_name "op ="}, _)) => rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
  | (_, Const (_, T')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
      in 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
        if rty = T' then rtrm
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   452
        else mk_repabs ctxt (rty, T') rtrm
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
      end   
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
  
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   455
  | _ => raise (LIFT_MATCH "injection (default)")
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
779
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   457
fun inj_repabs_trm_chk ctxt (rtrm, qtrm) =
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   458
  inj_repabs_trm ctxt (rtrm, qtrm) 
3b21b24a5fb6 corrected map declarations for Sum and Prod; moved absrep_fun examples in separate file
Christian Urban <urbanc@in.tum.de>
parents: 776
diff changeset
   459
  |> Syntax.check_term ctxt
776
d1064fa29424 renamed get_fun to absrep_fun; introduced explicit checked versions of the term functions
Christian Urban <urbanc@in.tum.de>
parents: 775
diff changeset
   460
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
end; (* structure *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464