src/Provers/trancl.ML
author berghofe
Mon, 22 Feb 2010 14:11:03 +0100
changeset 35281 206e2f1759cc
parent 33063 4d462963a7db
child 36692 54b64d4ad524
permissions -rw-r--r--
Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation contained variables bound by meta level quantifiers.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
     1
(*
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 26834
diff changeset
     2
    Title:      Transitivity reasoner for transitive closures of relations
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 26834
diff changeset
     3
    Author:     Oliver Kutter, TU Muenchen
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
     4
*)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
     5
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
     6
(*
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
     7
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
     8
The packages provides tactics trancl_tac and rtrancl_tac that prove
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
     9
goals of the form
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    10
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    11
   (x,y) : r^+     and     (x,y) : r^* (rtrancl_tac only)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    12
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    13
from premises of the form
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    14
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    15
   (x,y) : r,     (x,y) : r^+     and     (x,y) : r^* (rtrancl_tac only)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    16
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    17
by reflexivity and transitivity.  The relation r is determined by inspecting
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    18
the conclusion.
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    19
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    20
The package is implemented as an ML functor and thus not limited to
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    21
particular constructs for transitive and reflexive-transitive
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    22
closures, neither need relations be represented as sets of pairs.  In
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    23
order to instantiate the package for transitive closure only, supply
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    24
dummy theorems to the additional rules for reflexive-transitive
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    25
closures, and don't use rtrancl_tac!
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    26
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    27
*)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    28
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    29
signature TRANCL_ARITH =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    30
sig
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    31
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    32
  (* theorems for transitive closure *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    33
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    34
  val r_into_trancl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    35
      (* (a,b) : r ==> (a,b) : r^+ *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    36
  val trancl_trans : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    37
      (* [| (a,b) : r^+ ; (b,c) : r^+ |] ==> (a,c) : r^+ *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    38
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    39
  (* additional theorems for reflexive-transitive closure *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    40
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    41
  val rtrancl_refl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    42
      (* (a,a): r^* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    43
  val r_into_rtrancl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    44
      (* (a,b) : r ==> (a,b) : r^* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    45
  val trancl_into_rtrancl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    46
      (* (a,b) : r^+ ==> (a,b) : r^* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    47
  val rtrancl_trancl_trancl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    48
      (* [| (a,b) : r^* ; (b,c) : r^+ |] ==> (a,c) : r^+ *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    49
  val trancl_rtrancl_trancl : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    50
      (* [| (a,b) : r^+ ; (b,c) : r^* |] ==> (a,c) : r^+ *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    51
  val rtrancl_trans : thm
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    52
      (* [| (a,b) : r^* ; (b,c) : r^* |] ==> (a,c) : r^* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    53
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    54
  (* decomp: decompose a premise or conclusion
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    55
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    56
     Returns one of the following:
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    57
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
    58
     NONE if not an instance of a relation,
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
    59
     SOME (x, y, r, s) if instance of a relation, where
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    60
       x: left hand side argument, y: right hand side argument,
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    61
       r: the relation,
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    62
       s: the kind of closure, one of
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    63
            "r":   the relation itself,
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    64
            "r^+": transitive closure of the relation,
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    65
            "r^*": reflexive-transitive closure of the relation
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    66
  *)
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
    67
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    68
  val decomp: term ->  (term * term * term * string) option
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    69
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    70
end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    71
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    72
signature TRANCL_TAC =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    73
sig
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    74
  val trancl_tac: Proof.context -> int -> tactic
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    75
  val rtrancl_tac: Proof.context -> int -> tactic
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    76
end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    77
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    78
functor Trancl_Tac(Cls: TRANCL_ARITH): TRANCL_TAC =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    79
struct
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    80
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    81
22257
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    82
datatype proof
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    83
  = Asm of int
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    84
  | Thm of proof list * thm;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    85
22257
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    86
exception Cannot; (* internal exception: raised if no proof can be found *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    87
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
    88
fun decomp t = Option.map (fn (x, y, rel, r) =>
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
    89
  (Envir.beta_eta_contract x, Envir.beta_eta_contract y,
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
    90
   Envir.beta_eta_contract rel, r)) (Cls.decomp t);
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
    91
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
    92
fun prove thy r asms =
22257
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    93
  let
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    94
    fun inst thm =
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
    95
      let val SOME (_, _, r', _) = decomp (concl_of thm)
22257
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    96
      in Drule.cterm_instantiate [(cterm_of thy r', cterm_of thy r)] thm end;
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    97
    fun pr (Asm i) = List.nth (asms, i)
159bfab776e2 "prove" function now instantiates relation variable in order
berghofe
parents: 15574
diff changeset
    98
      | pr (Thm (prfs, thm)) = map pr prfs MRS inst thm
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
    99
  in pr end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   100
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   101
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   102
(* Internal datatype for inequalities *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   103
datatype rel
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   104
   = Trans  of term * term * proof  (* R^+ *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   105
   | RTrans of term * term * proof; (* R^* *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   106
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   107
 (* Misc functions for datatype rel *)
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   108
fun lower (Trans (x, _, _)) = x
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   109
  | lower (RTrans (x,_,_)) = x;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   110
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   111
fun upper (Trans (_, y, _)) = y
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   112
  | upper (RTrans (_,y,_)) = y;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   113
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   114
fun getprf   (Trans   (_, _, p)) = p
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   115
|   getprf   (RTrans (_,_, p)) = p;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   116
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   117
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   118
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   119
(*  mkasm_trancl Rel (t,n): term -> (term , int) -> rel list                *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   120
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   121
(*  Analyse assumption t with index n with respect to relation Rel:         *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   122
(*  If t is of the form "(x, y) : Rel" (or Rel^+), translate to             *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   123
(*  an object (singleton list) of internal datatype rel.                    *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   124
(*  Otherwise return empty list.                                            *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   125
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   126
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   127
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   128
fun mkasm_trancl  Rel  (t, n) =
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
   129
  case decomp t of
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   130
    SOME (x, y, rel,r) => if rel aconv Rel then
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   131
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   132
    (case r of
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   133
      "r"   => [Trans (x,y, Thm([Asm n], Cls.r_into_trancl))]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   134
    | "r+"  => [Trans (x,y, Asm n)]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   135
    | "r*"  => []
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   136
    | _     => error ("trancl_tac: unknown relation symbol"))
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   137
    else []
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   138
  | NONE => [];
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   139
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   140
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   141
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   142
(*  mkasm_rtrancl Rel (t,n): term -> (term , int) -> rel list               *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   143
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   144
(*  Analyse assumption t with index n with respect to relation Rel:         *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   145
(*  If t is of the form "(x, y) : Rel" (or Rel^+ or Rel^* ), translate to   *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   146
(*  an object (singleton list) of internal datatype rel.                    *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   147
(*  Otherwise return empty list.                                            *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   148
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   149
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   150
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   151
fun mkasm_rtrancl Rel (t, n) =
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
   152
  case decomp t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   153
   SOME (x, y, rel, r) => if rel aconv Rel then
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   154
    (case r of
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   155
      "r"   => [ Trans (x,y, Thm([Asm n], Cls.r_into_trancl))]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   156
    | "r+"  => [ Trans (x,y, Asm n)]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   157
    | "r*"  => [ RTrans(x,y, Asm n)]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   158
    | _     => error ("rtrancl_tac: unknown relation symbol" ))
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   159
   else []
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   160
  | NONE => [];
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   161
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   162
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   163
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   164
(*  mkconcl_trancl t: term -> (term, rel, proof)                            *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   165
(*  mkconcl_rtrancl t: term -> (term, rel, proof)                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   166
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   167
(*  Analyse conclusion t:                                                   *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   168
(*    - must be of form "(x, y) : r^+ (or r^* for rtrancl)                  *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   169
(*    - returns r                                                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   170
(*    - conclusion in internal form                                         *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   171
(*    - proof object                                                        *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   172
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   173
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   174
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   175
fun mkconcl_trancl  t =
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
   176
  case decomp t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   177
    SOME (x, y, rel, r) => (case r of
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   178
      "r+"  => (rel, Trans (x,y, Asm ~1), Asm 0)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   179
    | _     => raise Cannot)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   180
  | NONE => raise Cannot;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   181
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   182
fun mkconcl_rtrancl  t =
26834
87a5b9ec3863 Terms returned by decomp are now eta-contracted.
berghofe
parents: 22257
diff changeset
   183
  case decomp t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   184
    SOME (x,  y, rel,r ) => (case r of
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   185
      "r+"  => (rel, Trans (x,y, Asm ~1),  Asm 0)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   186
    | "r*"  => (rel, RTrans (x,y, Asm ~1), Asm 0)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   187
    | _     => raise Cannot)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15098
diff changeset
   188
  | NONE => raise Cannot;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   189
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   190
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   191
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   192
(*  makeStep (r1, r2): rel * rel -> rel                                     *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   193
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   194
(*  Apply transitivity to r1 and r2, obtaining a new element of r^+ or r^*, *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   195
(*  according the following rules:                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   196
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   197
(* ( (a, b) : r^+ , (b,c) : r^+ ) --> (a,c) : r^+                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   198
(* ( (a, b) : r^* , (b,c) : r^+ ) --> (a,c) : r^+                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   199
(* ( (a, b) : r^+ , (b,c) : r^* ) --> (a,c) : r^+                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   200
(* ( (a, b) : r^* , (b,c) : r^* ) --> (a,c) : r^*                           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   201
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   202
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   203
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   204
fun makeStep (Trans (a,_,p), Trans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.trancl_trans))
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   205
(* refl. + trans. cls. rules *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   206
|   makeStep (RTrans (a,_,p), Trans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.rtrancl_trancl_trancl))
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   207
|   makeStep (Trans (a,_,p), RTrans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.trancl_rtrancl_trancl))
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   208
|   makeStep (RTrans (a,_,p), RTrans(_,c,q))  = RTrans (a,c, Thm ([p,q], Cls.rtrancl_trans));
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   209
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   210
(* ******************************************************************* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   211
(*                                                                     *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   212
(* transPath (Clslist, Cls): (rel  list * rel) -> rel                  *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   213
(*                                                                     *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   214
(* If a path represented by a list of elements of type rel is found,   *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   215
(* this needs to be contracted to a single element of type rel.        *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   216
(* Prior to each transitivity step it is checked whether the step is   *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   217
(* valid.                                                              *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   218
(*                                                                     *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   219
(* ******************************************************************* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   220
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   221
fun transPath ([],acc) = acc
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   222
|   transPath (x::xs,acc) = transPath (xs, makeStep(acc,x))
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   223
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   224
(* ********************************************************************* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   225
(* Graph functions                                                       *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   226
(* ********************************************************************* *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   227
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   228
(* *********************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   229
(* Functions for constructing graphs                           *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   230
(* *********************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   231
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   232
fun addEdge (v,d,[]) = [(v,d)]
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   233
|   addEdge (v,d,((u,dl)::el)) = if v aconv u then ((v,d@dl)::el)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   234
    else (u,dl):: (addEdge(v,d,el));
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   235
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   236
(* ********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   237
(*                                                                        *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   238
(* mkGraph constructs from a list of objects of type rel  a graph g       *)
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   239
(* and a list of all edges with label r+.                                 *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   240
(*                                                                        *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   241
(* ********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   242
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   243
fun mkGraph [] = ([],[])
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   244
|   mkGraph ys =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   245
 let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   246
  fun buildGraph ([],g,zs) = (g,zs)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   247
  |   buildGraph (x::xs, g, zs) =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   248
        case x of (Trans (_,_,_)) =>
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   249
               buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), x::zs)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   250
        | _ => buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), zs)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   251
in buildGraph (ys, [], []) end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   252
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   253
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   254
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   255
(* adjacent g u : (''a * 'b list ) list -> ''a -> 'b list                  *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   256
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   257
(* List of successors of u in graph g                                      *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   258
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   259
(* *********************************************************************** *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   260
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   261
fun adjacent eq_comp ((v,adj)::el) u =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   262
    if eq_comp (u, v) then adj else adjacent eq_comp el u
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   263
|   adjacent _  []  _ = []
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   264
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   265
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   266
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   267
(* dfs eq_comp g u v:                                                      *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   268
(* ('a * 'a -> bool) -> ('a  *( 'a * rel) list) list ->                    *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   269
(* 'a -> 'a -> (bool * ('a * rel) list)                                    *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   270
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   271
(* Depth first search of v from u.                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   272
(* Returns (true, path(u, v)) if successful, otherwise (false, []).        *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   273
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   274
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   275
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   276
fun dfs eq_comp g u v =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   277
 let
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32285
diff changeset
   278
    val pred = Unsynchronized.ref [];
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32285
diff changeset
   279
    val visited = Unsynchronized.ref [];
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   280
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   281
    fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   282
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   283
    fun dfs_visit u' =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   284
    let val _ = visited := u' :: (!visited)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   285
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   286
    fun update (x,l) = let val _ = pred := (x,l) ::(!pred) in () end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   287
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   288
    in if been_visited v then ()
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   289
    else (app (fn (v',l) => if been_visited v' then () else (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   290
       update (v',l);
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   291
       dfs_visit v'; ()) )) (adjacent eq_comp g u')
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   292
     end
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   293
  in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   294
    dfs_visit u;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   295
    if (been_visited v) then (true, (!pred)) else (false , [])
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   296
  end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   297
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   298
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   299
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   300
(* transpose g:                                                            *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   301
(* (''a * ''a list) list -> (''a * ''a list) list                          *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   302
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   303
(* Computes transposed graph g' from g                                     *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   304
(* by reversing all edges u -> v to v -> u                                 *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   305
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   306
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   307
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   308
fun transpose eq_comp g =
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   309
  let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   310
   (* Compute list of reversed edges for each adjacency list *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   311
   fun flip (u,(v,l)::el) = (v,(u,l)) :: flip (u,el)
32768
wenzelm
parents: 32740
diff changeset
   312
     | flip (_,[]) = []
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   313
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   314
   (* Compute adjacency list for node u from the list of edges
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   315
      and return a likewise reduced list of edges.  The list of edges
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   316
      is searches for edges starting from u, and these edges are removed. *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   317
   fun gather (u,(v,w)::el) =
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   318
    let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   319
     val (adj,edges) = gather (u,el)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   320
    in
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   321
     if eq_comp (u, v) then (w::adj,edges)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   322
     else (adj,(v,w)::edges)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   323
    end
32768
wenzelm
parents: 32740
diff changeset
   324
   | gather (_,[]) = ([],[])
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   325
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   326
   (* For every node in the input graph, call gather to find all reachable
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   327
      nodes in the list of edges *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   328
   fun assemble ((u,_)::el) edges =
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   329
       let val (adj,edges) = gather (u,edges)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   330
       in (u,adj) :: assemble el edges
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   331
       end
32768
wenzelm
parents: 32740
diff changeset
   332
     | assemble [] _ = []
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   333
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   334
   (* Compute, for each adjacency list, the list with reversed edges,
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   335
      and concatenate these lists. *)
32768
wenzelm
parents: 32740
diff changeset
   336
   val flipped = maps flip g
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   337
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   338
 in assemble g flipped end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   339
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   340
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   341
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   342
(* dfs_reachable eq_comp g u:                                              *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   343
(* (int * int list) list -> int -> int list                                *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   344
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   345
(* Computes list of all nodes reachable from u in g.                       *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   346
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   347
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   348
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   349
fun dfs_reachable eq_comp g u =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   350
 let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   351
  (* List of vertices which have been visited. *)
32768
wenzelm
parents: 32740
diff changeset
   352
  val visited  = Unsynchronized.ref [];
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   353
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   354
  fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   355
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   356
  fun dfs_visit g u  =
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   357
      let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   358
   val _ = visited := u :: !visited
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   359
   val descendents =
30190
479806475f3c use long names for old-style fold combinators;
wenzelm
parents: 26834
diff changeset
   360
       List.foldr (fn ((v,l),ds) => if been_visited v then ds
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   361
            else v :: dfs_visit g v @ ds)
32768
wenzelm
parents: 32740
diff changeset
   362
        [] (adjacent eq_comp g u)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   363
   in  descendents end
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   364
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   365
 in u :: dfs_visit g u end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   366
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   367
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   368
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   369
(* dfs_term_reachable g u:                                                  *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   370
(* (term * term list) list -> term -> term list                            *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   371
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   372
(* Computes list of all nodes reachable from u in g.                       *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   373
(*                                                                         *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   374
(* *********************************************************************** *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   375
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   376
fun dfs_term_reachable g u = dfs_reachable (op aconv) g u;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   377
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   378
(* ************************************************************************ *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   379
(*                                                                          *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   380
(* findPath x y g: Term.term -> Term.term ->                                *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   381
(*                  (Term.term * (Term.term * rel list) list) ->            *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   382
(*                  (bool, rel list)                                        *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   383
(*                                                                          *)
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   384
(*  Searches a path from vertex x to vertex y in Graph g, returns true and  *)
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   385
(*  the list of edges if path is found, otherwise false and nil.            *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   386
(*                                                                          *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   387
(* ************************************************************************ *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   388
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   389
fun findPath x y g =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   390
  let
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   391
   val (found, tmp) =  dfs (op aconv) g x y ;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   392
   val pred = map snd tmp;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   393
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   394
   fun path x y  =
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   395
    let
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   396
         (* find predecessor u of node v and the edge u -> v *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   397
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   398
      fun lookup v [] = raise Cannot
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   399
      |   lookup v (e::es) = if (upper e) aconv v then e else lookup v es;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   400
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   401
      (* traverse path backwards and return list of visited edges *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   402
      fun rev_path v =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   403
        let val l = lookup v pred
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   404
            val u = lower l;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   405
        in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   406
          if u aconv x then [l] else (rev_path u) @ [l]
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   407
        end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   408
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   409
    in rev_path y end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   410
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   411
   in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   412
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   413
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   414
      if found then ( (found, (path x y) )) else (found,[])
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   415
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   416
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   417
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   418
   end;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   419
15098
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   420
(* ************************************************************************ *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   421
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   422
(* findRtranclProof g tranclEdges subgoal:                                  *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   423
(* (Term.term * (Term.term * rel list) list) -> rel -> proof list           *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   424
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   425
(* Searches in graph g a proof for subgoal.                                 *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   426
(*                                                                          *)
0726e7b15618 Documentation added/improved.
ballarin
parents: 15078
diff changeset
   427
(* ************************************************************************ *)
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   428
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   429
fun findRtranclProof g tranclEdges subgoal =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   430
   case subgoal of (RTrans (x,y,_)) => if x aconv y then [Thm ([], Cls.rtrancl_refl)] else (
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   431
     let val (found, path) = findPath (lower subgoal) (upper subgoal) g
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   432
     in
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   433
       if found then (
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   434
          let val path' = (transPath (tl path, hd path))
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   435
          in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   436
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   437
            case path' of (Trans (_,_,p)) => [Thm ([p], Cls.trancl_into_rtrancl )]
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   438
            | _ => [getprf path']
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   439
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   440
          end
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   441
       )
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   442
       else raise Cannot
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   443
     end
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   444
   )
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   445
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   446
| (Trans (x,y,_)) => (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   447
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   448
  let
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   449
   val Vx = dfs_term_reachable g x;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   450
   val g' = transpose (op aconv) g;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   451
   val Vy = dfs_term_reachable g' y;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   452
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   453
   fun processTranclEdges [] = raise Cannot
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   454
   |   processTranclEdges (e::es) =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   455
          if (upper e) mem Vx andalso (lower e) mem Vx
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   456
          andalso (upper e) mem Vy andalso (lower e) mem Vy
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   457
          then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   458
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   459
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   460
            if (lower e) aconv x then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   461
              if (upper e) aconv y then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   462
                  [(getprf e)]
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   463
              )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   464
              else (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   465
                  let
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   466
                    val (found,path) = findPath (upper e) y g
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   467
                  in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   468
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   469
                   if found then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   470
                       [getprf (transPath (path, e))]
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   471
                      ) else processTranclEdges es
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   472
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   473
                  end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   474
              )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   475
            )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   476
            else if (upper e) aconv y then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   477
               let val (xufound,xupath) = findPath x (lower e) g
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   478
               in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   479
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   480
                  if xufound then (
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   481
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   482
                    let val xuRTranclEdge = transPath (tl xupath, hd xupath)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   483
                            val xyTranclEdge = makeStep(xuRTranclEdge,e)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   484
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   485
                                in [getprf xyTranclEdge] end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   486
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   487
                 ) else processTranclEdges es
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   488
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   489
               end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   490
            )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   491
            else (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   492
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   493
                let val (xufound,xupath) = findPath x (lower e) g
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   494
                    val (vyfound,vypath) = findPath (upper e) y g
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   495
                 in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   496
                    if xufound then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   497
                         if vyfound then (
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   498
                            let val xuRTranclEdge = transPath (tl xupath, hd xupath)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   499
                                val vyRTranclEdge = transPath (tl vypath, hd vypath)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   500
                                val xyTranclEdge = makeStep (makeStep(xuRTranclEdge,e),vyRTranclEdge)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   501
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   502
                                in [getprf xyTranclEdge] end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   503
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   504
                         ) else processTranclEdges es
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   505
                    )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   506
                    else processTranclEdges es
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   507
                 end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   508
            )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   509
          )
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   510
          else processTranclEdges es;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   511
   in processTranclEdges tranclEdges end )
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   512
| _ => raise Cannot
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   513
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   514
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   515
fun solveTrancl (asms, concl) =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   516
 let val (g,_) = mkGraph asms
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   517
 in
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   518
  let val (_, subgoal, _) = mkconcl_trancl concl
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   519
      val (found, path) = findPath (lower subgoal) (upper subgoal) g
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   520
  in
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   521
    if found then  [getprf (transPath (tl path, hd path))]
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   522
    else raise Cannot
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   523
  end
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   524
 end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   525
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   526
fun solveRtrancl (asms, concl) =
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   527
 let val (g,tranclEdges) = mkGraph asms
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   528
     val (_, subgoal, _) = mkconcl_rtrancl concl
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   529
in
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   530
  findRtranclProof g tranclEdges subgoal
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   531
end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   532
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   533
32277
ff1e59a15146 trans_tac: use theory from goal state, not the static context, which seems to be outdated under certain circumstances (why?);
wenzelm
parents: 32215
diff changeset
   534
fun trancl_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   535
 let
32285
ab9b66c2bbca trancl_tac etc.: back to static context -- problem was caused by bad solver in AFP/JiveDataStoreModel;
wenzelm
parents: 32283
diff changeset
   536
  val thy = ProofContext.theory_of ctxt;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   537
  val Hs = Logic.strip_assums_hyp A;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   538
  val C = Logic.strip_assums_concl A;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   539
  val (rel, subgoals, prf) = mkconcl_trancl C;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   540
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32768
diff changeset
   541
  val prems = flat (map_index (mkasm_trancl rel o swap) Hs);
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   542
  val prfs = solveTrancl (prems, C);
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   543
 in
35281
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   544
  Subgoal.FOCUS (fn {prems, concl, ...} =>
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   545
    let
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   546
      val SOME (_, _, rel', _) = decomp (term_of concl);
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   547
      val thms = map (prove thy rel' prems) prfs
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   548
    in rtac (prove thy rel' thms prf) 1 end) ctxt n st
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   549
 end
32277
ff1e59a15146 trans_tac: use theory from goal state, not the static context, which seems to be outdated under certain circumstances (why?);
wenzelm
parents: 32215
diff changeset
   550
 handle Cannot => Seq.empty);
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   551
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   552
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   553
fun rtrancl_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   554
 let
32285
ab9b66c2bbca trancl_tac etc.: back to static context -- problem was caused by bad solver in AFP/JiveDataStoreModel;
wenzelm
parents: 32283
diff changeset
   555
  val thy = ProofContext.theory_of ctxt;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   556
  val Hs = Logic.strip_assums_hyp A;
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   557
  val C = Logic.strip_assums_concl A;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   558
  val (rel, subgoals, prf) = mkconcl_rtrancl C;
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   559
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32768
diff changeset
   560
  val prems = flat (map_index (mkasm_rtrancl rel o swap) Hs);
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   561
  val prfs = solveRtrancl (prems, C);
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   562
 in
35281
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   563
  Subgoal.FOCUS (fn {prems, concl, ...} =>
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   564
    let
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   565
      val SOME (_, _, rel', _) = decomp (term_of concl);
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   566
      val thms = map (prove thy rel' prems) prfs
206e2f1759cc Fixed bug that caused (r)trancl_tac to crash when the term denoting the relation
berghofe
parents: 33063
diff changeset
   567
    in rtac (prove thy rel' thms prf) 1 end) ctxt n st
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   568
 end
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 30190
diff changeset
   569
 handle Cannot => Seq.empty | Subscript => Seq.empty);
15076
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   570
4b3d280ef06a New prover for transitive and reflexive-transitive closure of relations.
ballarin
parents:
diff changeset
   571
end;