Quot/quotient_term.ML
author Christian Urban <urbanc@in.tum.de>
Tue, 22 Dec 2009 21:16:11 +0100
changeset 775 26fefde1d124
parent 774 b4ffb8826105
child 776 d1064fa29424
permissions -rw-r--r--
tuned
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
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
     6
   val get_fun: 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
     7
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
     8
   val regularize_trm: Proof.context -> (term * term) -> 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
   val inj_repabs_trm: Proof.context -> (term * term) -> term
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
end;
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
structure Quotient_Term: QUOTIENT_TERM =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
struct
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    15
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
    16
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
    17
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
    18
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
    19
(* Calculates the aggregate abs and rep functions for a given type; *) 
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
    20
(* repF is for constants' arguments; absF is for constants;         *)
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
(* function types need to be treated specially, since repF and 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
    22
(* change                                                           *)
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
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
    24
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
    25
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
    26
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
    27
  | 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
    28
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
    29
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
    30
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
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
    32
  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
    33
    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
    34
  | 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
    35
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
fun get_fun_aux lthy s fs =
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
    37
let
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
  val thy = ProofContext.theory_of lthy
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
  val exc = LIFT_MATCH (space_implode " " ["get_fun_aux: no map for type", quote s, "."])
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
    40
  val info = maps_lookup thy s handle NotFound => raise exc
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
    41
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
    42
  list_comb (Const (#mapfun info, dummyT), fs)
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
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
    44
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
fun get_const flag lthy _ qty =
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
(* FIXME: check here that the type-constructors of _ and qty are related *)
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
let
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
    48
  val thy = ProofContext.theory_of lthy
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
  val qty_name = Long_Name.base_name (fst (dest_Type qty))
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
    50
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
    51
  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
    52
    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
    53
  | 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
    54
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
    55
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
fun get_fun flag lthy (rty, qty) =
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
    57
  if rty = qty then mk_identity qty else
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
    58
  case (rty, qty) 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
    59
    (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
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
    60
     let
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
    61
       val fs_ty1 = get_fun (negF flag) lthy (ty1, ty1')
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
    62
       val fs_ty2 = get_fun flag lthy (ty2, ty2')
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
    63
     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
    64
       get_fun_aux lthy "fun" [fs_ty1, fs_ty2]
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
    65
     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
    66
  | (Type (s, _), Type (s', [])) =>
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
    67
     if s = s'
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
    68
     then mk_identity qty
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
    69
     else get_const flag lthy rty qty
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
    70
  | (Type (s, tys), Type (s', tys')) =>
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
    71
     let
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
    72
        val args = map (get_fun flag lthy) (tys ~~ tys')
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
    73
     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
    74
        if s = s'
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
    75
        then get_fun_aux lthy s args
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
    76
        else mk_compose flag (get_const flag lthy rty qty, get_fun_aux lthy s args)
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
    77
     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
    78
  | (TFree x, TFree x') =>
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
    79
     if x = x'
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
    80
     then mk_identity qty
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
    81
     else raise (LIFT_MATCH "get_fun (frees)")
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
    82
  | (TVar _, TVar _) => raise (LIFT_MATCH "get_fun (vars)")
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
    83
  | _ => raise (LIFT_MATCH "get_fun (default)")
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
    84
762
baac4639ecef avoided global "open"s - replaced by local "open"s
Christian Urban <urbanc@in.tum.de>
parents: 761
diff changeset
    85
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
(*
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
Regularizing an rtrm means:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
 
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    89
 - 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
    90
   by bounded quantifiers, for example:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    92
      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
    93
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    94
   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
    95
 
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    96
 - 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
    97
   by bounded abstractions, for example:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
      
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
    99
      %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
   100
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   101
 - 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
   102
   corresponding equivalence relations, for example:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   104
      A = B  ----> R A B
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
   or 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   108
      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
   109
 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
   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
   111
*)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   113
(* 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
   114
(* that will be the argument of Respects     *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
fun mk_resp_arg lthy (rty, qty) =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  val thy = ProofContext.theory_of lthy
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
in  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  if rty = qty
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  then HOLogic.eq_const rty
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  else
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
    case (rty, qty) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
      (Type (s, tys), Type (s', tys')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
       if s = s' 
761
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   125
       then 
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   126
         let
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   127
           val exc = LIFT_MATCH ("mk_resp_arg (no map function found for type " ^ s) 
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   128
           val map_info = maps_lookup thy s handle NotFound => raise exc
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   129
           val args = map (mk_resp_arg lthy) (tys ~~ tys')
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   130
         in
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   131
           list_comb (Const (#relfun map_info, dummyT), args) 
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   132
         end  
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   133
       else 
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   134
         let  
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   135
           val SOME qinfo = quotdata_lookup_thy thy s'
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   136
           (* FIXME: check in this case that the rty and qty *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   137
           (* FIXME: correspond to each other *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   138
           val (s, _) = dest_Const (#rel qinfo)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   139
           (* FIXME: the relation should only be the string        *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   140
           (* FIXME: and the type needs to be calculated as below; *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   141
           (* FIXME: maybe one should actually have a term         *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   142
           (* FIXME: and one needs to force it to have this type   *)
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   143
         in
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   144
           Const (s, rty --> rty --> @{typ bool})
e2ac18492c68 small tuning
Christian Urban <urbanc@in.tum.de>
parents: 760
diff changeset
   145
         end
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
      | _ => HOLogic.eq_const dummyT 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
             (* 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
   148
end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
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
   151
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
   152
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
   153
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
   154
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
(* - 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
   156
(*   otherwise to the given term,                  *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
(* - used by regularize, therefore abstracted      *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
(*   variables do not have to be treated specially *)
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   159
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
   160
  case (trm1, trm2) of
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   161
    (Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t'))
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   162
  | _ => f (trm1, trm2)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
(* 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
   165
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
   166
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
(* produces a regularized version of rtrm       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
(*                                              *)
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   170
(* - the result still contains dummyTs          *)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
(*                                              *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
(* - 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
   173
(*   special treatment of bound variables       *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   175
fun regularize_trm lthy (rtrm, qtrm) =
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  case (rtrm, qtrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
    (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
   178
       let
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   179
         val subtrm = Abs(x, ty, regularize_trm lthy (t, t'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
         if ty = ty' then subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
         else mk_babs $ (mk_resp $ mk_resp_arg lthy (ty, ty')) $ subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  | (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
   186
       let
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   187
         val subtrm = apply_subt (regularize_trm lthy) (t, t')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
         if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
         else mk_ball $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  | (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
   194
       let
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   195
         val subtrm = apply_subt (regularize_trm lthy) (t, t')
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
         if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
         else mk_bex $ (mk_resp $ mk_resp_arg lthy (qnt_typ ty, qnt_typ ty')) $ subtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
       end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  | (* 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
   202
    (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
   203
         if ty = ty' then rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
         else mk_resp_arg lthy (domain_type ty, domain_type ty') 
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
  | (* 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
   207
    (rel, Const (@{const_name "op ="}, ty')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
       let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
         val exc = LIFT_MATCH "regularise (relation mismatch)"
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
         val rel_ty = (fastype_of rel) handle TERM _ => raise exc 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
         val rel' = mk_resp_arg lthy (domain_type rel_ty, domain_type ty') 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
       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
   213
         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
   214
       end  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  | (_, Const _) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
       let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
         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
   219
           | same_name _ _ = false
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
          (* 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
   221
          (*             Why?                     *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
          (* 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
   223
             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
   224
             of different types they do differ.
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
        
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
             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
   227
             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
   228
    
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
             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
   230
             cleaning. *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
           (* 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
   232
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
         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
   234
         else 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
           let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
             val thy = ProofContext.theory_of lthy
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
             val qtrm_str = Syntax.string_of_term lthy qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
             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
   239
             val exc2 = LIFT_MATCH ("regularize (constant " ^ qtrm_str ^ " mismatch)")
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
             val rtrm' = (#rconst (qconsts_lookup thy qtrm)) handle NotFound => raise exc1
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
           in 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
             if Pattern.matches thy (rtrm', rtrm) 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
             then rtrm else raise exc2
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
           end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
       end 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
  | (t1 $ t2, t1' $ t2') =>
775
Christian Urban <urbanc@in.tum.de>
parents: 774
diff changeset
   248
       (regularize_trm lthy (t1, t1')) $ (regularize_trm lthy (t2, t2'))
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  | (Bound i, Bound i') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
       if i = i' then rtrm 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
       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
   253
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  | _ =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
       let 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
         val rtrm_str = Syntax.string_of_term lthy rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
         val qtrm_str = Syntax.string_of_term lthy qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
       in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
         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
   260
       end
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
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
(*
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   265
Injection of Rep/Abs means:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  For abstractions
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
:
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   269
  * 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
   270
    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
   271
 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
  For applications:
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   274
  * 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
   275
    the second argument. If the application is a bounded abstraction,
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   276
    we always put an Re/Abs around it (since bounded abstractions
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   277
    always need lifting). Otherwise we recurse on both arguments.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   279
  For constants:
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   281
  * 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
   282
    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
   283
    and Rep/Abs around it. 
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   285
  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
   286
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   287
  * 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
   288
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   289
  Vars case cannot occur.
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
*)
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
fun mk_repabs lthy (T, T') trm = 
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
   293
  get_fun repF lthy (T, T') $ (get_fun absF lthy (T, T') $ trm)
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
(* 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
   297
(* 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
   298
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
fun inj_repabs_trm lthy (rtrm, qtrm) =
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
 case (rtrm, qtrm) of
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
       Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm lthy (t, t'))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
       Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm lthy (t, t'))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  | (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
   308
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
        val qty = fastype_of qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
      in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
        mk_repabs lthy (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm lthy (t, t')))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
      end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
  | (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
   316
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
        val qty = fastype_of qtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
        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
   320
        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
   321
        val yvar = Free (y, T)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
        val result = Term.lambda_name (y, yvar) (inj_repabs_trm lthy (s, s'))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
      in
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
        if rty = qty then result
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
        else mk_repabs lthy (rty, qty) result
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
      end
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
  | (t $ s, t' $ s') =>  
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
       (inj_repabs_trm lthy (t, t')) $ (inj_repabs_trm lthy (s, s'))
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  | (Free (_, T), Free (_, T')) => 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
        if T = T' then rtrm 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
        else mk_repabs lthy (T, T') rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  | (_, Const (@{const_name "op ="}, _)) => rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  | (_, Const (_, T')) =>
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
      let
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
        val rty = fastype_of rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
      in 
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
        if rty = T' then rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
        else mk_repabs lthy (rty, T') rtrm
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
      end   
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
  
760
c1989de100b4 various tunings; map_lookup now raises an exception; addition to FIXME-TODO
Christian Urban <urbanc@in.tum.de>
parents: 758
diff changeset
   345
  | _ => raise (LIFT_MATCH "injection (default)")
758
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
end; (* structure *)
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
3104d62e7a16 moved the QuotMain code into two ML-files
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350