src/Provers/quasi.ML
author wenzelm
Fri, 23 Nov 2018 16:43:11 +0100
changeset 69334 6b49700da068
parent 67379 c2dfc510a38c
permissions -rw-r--r--
clarified font_domain: strict excludes e.g. space character;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37744
3daaf23b9ab4 tuned titles
haftmann
parents: 33063
diff changeset
     1
(*  Title:      Provers/quasi.ML
3daaf23b9ab4 tuned titles
haftmann
parents: 33063
diff changeset
     2
    Author:     Oliver Kutter, TU Muenchen
29276
94b1ffec9201 qualified Term.rename_wrt_term;
wenzelm
parents: 22578
diff changeset
     3
94b1ffec9201 qualified Term.rename_wrt_term;
wenzelm
parents: 22578
diff changeset
     4
Reasoner for simple transitivity and quasi orders.
94b1ffec9201 qualified Term.rename_wrt_term;
wenzelm
parents: 22578
diff changeset
     5
*)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
     6
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
     7
(*
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
     8
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
     9
The package provides tactics trans_tac and quasi_tac that use
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    10
premises of the form
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    11
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    12
  t = u, t ~= u, t < u and t <= u
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    13
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    14
to
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    15
- either derive a contradiction, in which case the conclusion can be
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    16
  any term,
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    17
- or prove the concluson, which must be of the form t ~= u, t < u or
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    18
  t <= u.
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    19
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    20
Details:
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    21
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    22
1. trans_tac:
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    23
   Only premises of form t <= u are used and the conclusion must be of
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    24
   the same form.  The conclusion is proved, if possible, by a chain of
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    25
   transitivity from the assumptions.
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    26
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    27
2. quasi_tac:
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    28
   <= is assumed to be a quasi order and < its strict relative, defined
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    29
   as t < u == t <= u & t ~= u.  Again, the conclusion is proved from
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    30
   the assumptions.
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    31
   Note that the presence of a strict relation is not necessary for
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    32
   quasi_tac.  Configure decomp_quasi to ignore < and ~=.  A list of
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    33
   required theorems for both situations is given below.
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    34
*)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    35
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    36
signature LESS_ARITH =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    37
sig
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    38
  (* Transitivity of <=
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    39
     Note that transitivities for < hold for partial orders only. *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    40
  val le_trans: thm  (* [| x <= y; y <= z |] ==> x <= z *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    41
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    42
  (* Additional theorem for quasi orders *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    43
  val le_refl: thm  (* x <= x *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    44
  val eqD1: thm (* x = y ==> x <= y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    45
  val eqD2: thm (* x = y ==> y <= x *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    46
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    47
  (* Additional theorems for premises of the form x < y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    48
  val less_reflE: thm  (* x < x ==> P *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    49
  val less_imp_le : thm (* x < y ==> x <= y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    50
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    51
  (* Additional theorems for premises of the form x ~= y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    52
  val le_neq_trans : thm (* [| x <= y ; x ~= y |] ==> x < y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    53
  val neq_le_trans : thm (* [| x ~= y ; x <= y |] ==> x < y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    54
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    55
  (* Additional theorem for goals of form x ~= y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    56
  val less_imp_neq : thm (* x < y ==> x ~= y *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    57
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    58
  (* Analysis of premises and conclusion *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
    59
  (* decomp_x (`x Rel y') should yield SOME (x, Rel, y)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    60
       where Rel is one of "<", "<=", "=" and "~=",
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    61
       other relation symbols cause an error message *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    62
  (* decomp_trans is used by trans_tac, it may only return Rel = "<=" *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
    63
  val decomp_trans: theory -> term -> (term * string * term) option
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    64
  (* decomp_quasi is used by quasi_tac *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
    65
  val decomp_quasi: theory -> term -> (term * string * term) option
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    66
end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    67
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    68
signature QUASI_TAC =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    69
sig
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    70
  val trans_tac: Proof.context -> int -> tactic
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    71
  val quasi_tac: Proof.context -> int -> tactic
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    72
end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    73
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    74
functor Quasi_Tac(Less: LESS_ARITH): QUASI_TAC =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    75
struct
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    76
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    77
(* Internal datatype for the proof *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    78
datatype proof
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    79
  = Asm of int
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    80
  | Thm of proof list * thm;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    81
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    82
exception Cannot;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    83
 (* Internal exception, raised if conclusion cannot be derived from
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    84
     assumptions. *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    85
exception Contr of proof;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    86
  (* Internal exception, raised if contradiction ( x < x ) was derived *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    87
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    88
fun prove asms =
42364
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
    89
  let
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
    90
    fun pr (Asm i) = nth asms i
8c674b3b8e44 eliminated old List.nth;
wenzelm
parents: 42361
diff changeset
    91
      | pr (Thm (prfs, thm)) = map pr prfs MRS thm;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    92
  in pr end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    93
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    94
(* Internal datatype for inequalities *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    95
datatype less
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    96
   = Less  of term * term * proof
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    97
   | Le    of term * term * proof
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
    98
   | NotEq of term * term * proof;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
    99
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   100
 (* Misc functions for datatype less *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   101
fun lower (Less (x, _, _)) = x
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   102
  | lower (Le (x, _, _)) = x
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   103
  | lower (NotEq (x,_,_)) = x;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   104
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   105
fun upper (Less (_, y, _)) = y
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   106
  | upper (Le (_, y, _)) = y
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   107
  | upper (NotEq (_,y,_)) = y;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   108
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   109
fun getprf   (Less (_, _, p)) = p
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   110
|   getprf   (Le   (_, _, p)) = p
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   111
|   getprf   (NotEq (_,_, p)) = p;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   112
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   113
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   114
(*                                                                          *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   115
(* mkasm_trans sign (t, n) :  theory -> (Term.term * int)  -> less          *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   116
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   117
(* Tuple (t, n) (t an assumption, n its index in the assumptions) is        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   118
(* translated to an element of type less.                                   *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   119
(* Only assumptions of form x <= y are used, all others are ignored         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   120
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   121
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   122
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   123
fun mkasm_trans thy (t, n) =
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   124
  case Less.decomp_trans thy t of
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   125
    SOME (x, rel, y) =>
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   126
    (case rel of
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   127
     "<="  =>  [Le (x, y, Asm n)]
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   128
    | _     => error ("trans_tac: unknown relation symbol ``" ^ rel ^
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   129
                 "''returned by decomp_trans."))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   130
  | NONE => [];
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   131
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   132
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   133
(*                                                                          *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   134
(* mkasm_quasi sign (t, n) : theory -> (Term.term * int) -> less            *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   135
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   136
(* Tuple (t, n) (t an assumption, n its index in the assumptions) is        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   137
(* translated to an element of type less.                                   *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   138
(* Quasi orders only.                                                       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   139
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   140
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   141
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   142
fun mkasm_quasi thy (t, n) =
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   143
  case Less.decomp_quasi thy t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   144
    SOME (x, rel, y) => (case rel of
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   145
      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE))
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   146
               else [Less (x, y, Asm n)]
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   147
    | "<="  => [Le (x, y, Asm n)]
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   148
    | "="   => [Le (x, y, Thm ([Asm n], Less.eqD1)),
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   149
                Le (y, x, Thm ([Asm n], Less.eqD2))]
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   150
    | "~="  => if (x aconv y) then
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   151
                  raise Contr (Thm ([(Thm ([(Thm ([], Less.le_refl)) ,(Asm n)], Less.le_neq_trans))], Less.less_reflE))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   152
               else [ NotEq (x, y, Asm n),
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   153
                      NotEq (y, x,Thm ( [Asm n], @{thm not_sym}))]
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   154
    | _     => error ("quasi_tac: unknown relation symbol ``" ^ rel ^
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   155
                 "''returned by decomp_quasi."))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   156
  | NONE => [];
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   157
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   158
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   159
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   160
(*                                                                          *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   161
(* mkconcl_trans sign t : theory -> Term.term -> less                       *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   162
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   163
(* Translates conclusion t to an element of type less.                      *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   164
(* Only for Conclusions of form x <= y or x < y.                            *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   165
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   166
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   167
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   168
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   169
fun mkconcl_trans thy t =
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   170
  case Less.decomp_trans thy t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   171
    SOME (x, rel, y) => (case rel of
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   172
     "<="  => (Le (x, y, Asm ~1), Asm 0)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   173
    | _  => raise Cannot)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   174
  | NONE => raise Cannot;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   175
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   176
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   177
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   178
(*                                                                          *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   179
(* mkconcl_quasi sign t : theory -> Term.term -> less                       *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   180
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   181
(* Translates conclusion t to an element of type less.                      *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   182
(* Quasi orders only.                                                       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   183
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   184
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   185
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   186
fun mkconcl_quasi thy t =
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   187
  case Less.decomp_quasi thy t of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   188
    SOME (x, rel, y) => (case rel of
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   189
      "<"   => ([Less (x, y, Asm ~1)], Asm 0)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   190
    | "<="  => ([Le (x, y, Asm ~1)], Asm 0)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   191
    | "~="  => ([NotEq (x,y, Asm ~1)], Asm 0)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   192
    | _  => raise Cannot)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   193
| NONE => raise Cannot;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   194
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   195
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   196
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   197
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   198
(* mergeLess (less1,less2):  less * less -> less                       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   199
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   200
(* Merge to elements of type less according to the following rules     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   201
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   202
(* x <= y && y <= z ==> x <= z                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   203
(* x <= y && x ~= y ==> x < y                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   204
(* x ~= y && x <= y ==> x < y                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   205
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   206
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   207
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   208
fun mergeLess (Le (x, _, p) , Le (_ , z, q)) =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   209
      Le (x, z, Thm ([p,q] , Less.le_trans))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   210
|   mergeLess (Le (x, z, p) , NotEq (x', z', q)) =
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   211
      if (x aconv x' andalso z aconv z' )
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   212
       then Less (x, z, Thm ([p,q] , Less.le_neq_trans))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   213
        else error "quasi_tac: internal error le_neq_trans"
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   214
|   mergeLess (NotEq (x, z, p) , Le (x' , z', q)) =
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   215
      if (x aconv x' andalso z aconv z')
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   216
      then Less (x, z, Thm ([p,q] , Less.neq_le_trans))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   217
      else error "quasi_tac: internal error neq_le_trans"
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   218
|   mergeLess (_, _) =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   219
      error "quasi_tac: internal error: undefined case";
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   220
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   221
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   222
(* ******************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   223
(* tr checks for valid transitivity step                                *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   224
(* ******************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   225
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   226
infix tr;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   227
fun (Le (_, y, _))   tr (Le (x', _, _))   = ( y aconv x' )
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   228
  | _ tr _ = false;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   229
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   230
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   231
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   232
(* transPath (Lesslist, Less): (less list * less) -> less              *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   233
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   234
(* If a path represented by a list of elements of type less is found,  *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   235
(* this needs to be contracted to a single element of type less.       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   236
(* Prior to each transitivity step it is checked whether the step is   *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   237
(* valid.                                                              *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   238
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   239
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   240
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   241
fun transPath ([],lesss) = lesss
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   242
|   transPath (x::xs,lesss) =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   243
      if lesss tr x then transPath (xs, mergeLess(lesss,x))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   244
      else error "trans/quasi_tac: internal error transpath";
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   245
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   246
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   247
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   248
(* less1 subsumes less2 : less -> less -> bool                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   249
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   250
(* subsumes checks whether less1 implies less2                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   251
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   252
(* ******************************************************************* *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   253
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   254
infix subsumes;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   255
fun (Le (x, y, _)) subsumes (Le (x', y', _)) =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   256
      (x aconv x' andalso y aconv y')
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   257
  | (Le _) subsumes (Less _) =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   258
      error "trans/quasi_tac: internal error: Le cannot subsume Less"
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   259
  | (NotEq(x,y,_)) subsumes (NotEq(x',y',_)) = x aconv x' andalso y aconv y' orelse x aconv y' andalso y aconv x'
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   260
  | _ subsumes _ = false;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   261
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   262
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   263
(*                                                                     *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   264
(* triv_solv less1 : less ->  proof option                     *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   265
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   266
(* Solves trivial goal x <= x.                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   267
(*                                                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   268
(* ******************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   269
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   270
fun triv_solv (Le (x, x', _)) =
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   271
    if x aconv x' then  SOME (Thm ([], Less.le_refl))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   272
    else NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   273
|   triv_solv _ = NONE;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   274
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   275
(* ********************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   276
(* Graph functions                                                       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   277
(* ********************************************************************* *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   278
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   279
(* *********************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   280
(* Functions for constructing graphs                           *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   281
(* *********************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   282
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   283
fun addEdge (v,d,[]) = [(v,d)]
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   284
|   addEdge (v,d,((u,dl)::el)) = if v aconv u then ((v,d@dl)::el)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   285
    else (u,dl):: (addEdge(v,d,el));
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   286
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   287
(* ********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   288
(*                                                                        *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   289
(* mkQuasiGraph constructs from a list of objects of type less a graph g, *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   290
(* by taking all edges that are candidate for a <=, and a list neqE, by   *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   291
(* taking all edges that are candiate for a ~=                            *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   292
(*                                                                        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   293
(* ********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   294
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   295
fun mkQuasiGraph [] = ([],[])
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   296
|   mkQuasiGraph lessList =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   297
 let
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   298
 fun buildGraphs ([],leG, neqE) = (leG,  neqE)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   299
  |   buildGraphs (l::ls, leG,  neqE) = case l of
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   300
       (Less (x,y,p)) =>
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   301
         let
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   302
          val leEdge  = Le (x,y, Thm ([p], Less.less_imp_le))
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   303
          val neqEdges = [ NotEq (x,y, Thm ([p], Less.less_imp_neq)),
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   304
                           NotEq (y,x, Thm ( [Thm ([p], Less.less_imp_neq)], @{thm not_sym}))]
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   305
         in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   306
           buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,leEdge)],leG))), neqEdges@neqE)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   307
         end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   308
     |  (Le (x,y,p))   => buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,l)],leG))), neqE)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   309
     | _ =>  buildGraphs (ls, leG,  l::neqE) ;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   310
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   311
in buildGraphs (lessList, [],  []) end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   312
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   313
(* ********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   314
(*                                                                        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   315
(* mkGraph constructs from a list of objects of type less a graph g       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   316
(* Used for plain transitivity chain reasoning.                           *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   317
(*                                                                        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   318
(* ********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   319
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   320
fun mkGraph [] = []
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   321
|   mkGraph lessList =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   322
 let
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   323
  fun buildGraph ([],g) = g
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   324
  |   buildGraph (l::ls, g) =  buildGraph (ls, (addEdge ((lower l),[((upper l),l)],g)))
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   325
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   326
in buildGraph (lessList, []) end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   327
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   328
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   329
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   330
(* adjacent g u : (''a * 'b list ) list -> ''a -> 'b list                  *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   331
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   332
(* List of successors of u in graph g                                      *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   333
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   334
(* *********************************************************************** *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   335
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   336
fun adjacent eq_comp ((v,adj)::el) u =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   337
    if eq_comp (u, v) then adj else adjacent eq_comp el u
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   338
|   adjacent _  []  _ = []
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   339
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   340
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   341
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   342
(* dfs eq_comp g u v:                                                      *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   343
(* ('a * 'a -> bool) -> ('a  *( 'a * less) list) list ->                   *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   344
(* 'a -> 'a -> (bool * ('a * less) list)                                   *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   345
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   346
(* Depth first search of v from u.                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   347
(* Returns (true, path(u, v)) if successful, otherwise (false, []).        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   348
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   349
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   350
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   351
fun dfs eq_comp g u v =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   352
 let
32740
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32285
diff changeset
   353
    val pred = Unsynchronized.ref [];
9dd0a2f83429 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32285
diff changeset
   354
    val visited = Unsynchronized.ref [];
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   355
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   356
    fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   357
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   358
    fun dfs_visit u' =
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   359
    let val _ = visited := u' :: (!visited)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   360
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   361
    fun update (x,l) = let val _ = pred := (x,l) ::(!pred) in () end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   362
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   363
    in if been_visited v then ()
67379
c2dfc510a38c prefer qualified names;
wenzelm
parents: 59584
diff changeset
   364
    else (List.app (fn (v',l) => if been_visited v' then () else (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   365
       update (v',l);
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   366
       dfs_visit v'; ()) )) (adjacent eq_comp g u')
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   367
     end
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   368
  in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   369
    dfs_visit u;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   370
    if (been_visited v) then (true, (!pred)) else (false , [])
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   371
  end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   372
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   373
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   374
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   375
(* Begin: Quasi Order relevant functions                                    *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   376
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   377
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   378
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   379
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   380
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   381
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   382
(* findPath x y g: Term.term -> Term.term ->                                *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   383
(*                  (Term.term * (Term.term * less list) list) ->           *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   384
(*                  (bool, less list)                                       *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   385
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   386
(*  Searches a path from vertex x to vertex y in Graph g, returns true and  *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   387
(*  the list of edges forming the path, if a path is found, otherwise false *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   388
(*  and nil.                                                                *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   389
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   390
(* ************************************************************************ *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   391
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   392
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   393
 fun findPath x y g =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   394
  let
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   395
    val (found, tmp) =  dfs (op aconv) g x y ;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   396
    val pred = map snd tmp;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   397
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   398
    fun path x y  =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   399
      let
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   400
       (* find predecessor u of node v and the edge u -> v *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   401
       fun lookup v [] = raise Cannot
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   402
       |   lookup v (e::es) = if (upper e) aconv v then e else lookup v es;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   403
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   404
       (* traverse path backwards and return list of visited edges *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   405
       fun rev_path v =
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   406
        let val l = lookup v pred
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   407
            val u = lower l;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   408
        in
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   409
           if u aconv x then [l] else (rev_path u) @ [l]
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   410
        end
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   411
      in rev_path y end;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   412
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   413
  in
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   414
   if found then (
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   415
    if x aconv y then (found,[(Le (x, y, (Thm ([], Less.le_refl))))])
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   416
    else (found, (path x y) ))
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   417
   else (found,[])
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   418
  end;
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   419
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   420
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   421
(* ************************************************************************ *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   422
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   423
(* findQuasiProof (leqG, neqE) subgoal:                                     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   424
(* (Term.term * (Term.term * less list) list) * less list  -> less -> proof *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   425
(*                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   426
(* Constructs a proof for subgoal by searching a special path in leqG and   *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   427
(* neqE. Raises Cannot if construction of the proof fails.                  *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   428
(*                                                                          *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   429
(* ************************************************************************ *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   430
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   431
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   432
(* As the conlusion can be either of form x <= y, x < y or x ~= y we have        *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   433
(* three cases to deal with. Finding a transitivity path from x to y with label  *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   434
(* 1. <=                                                                         *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   435
(*    This is simply done by searching any path from x to y in the graph leG.    *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   436
(*    The graph leG contains only edges with label <=.                           *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   437
(*                                                                               *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   438
(* 2. <                                                                          *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   439
(*    A path from x to y with label < can be found by searching a path with      *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   440
(*    label <= from x to y in the graph leG and merging the path x <= y with     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   441
(*    a parallel edge x ~= y resp. y ~= x to x < y.                              *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   442
(*                                                                               *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   443
(* 3. ~=                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   444
(*   If the conclusion is of form x ~= y, we can find a proof either directly,   *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   445
(*   if x ~= y or y ~= x are among the assumptions, or by constructing x ~= y if *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   446
(*   x < y or y < x follows from the assumptions.                                *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   447
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   448
fun findQuasiProof (leG, neqE) subgoal =
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   449
  case subgoal of (Le (x,y, _)) => (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   450
   let
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   451
    val (xyLefound,xyLePath) = findPath x y leG
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   452
   in
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   453
    if xyLefound then (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   454
     let
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   455
      val Le_x_y = (transPath (tl xyLePath, hd xyLePath))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   456
     in getprf Le_x_y end )
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   457
    else raise Cannot
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   458
   end )
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   459
  | (Less (x,y,_))  => (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   460
   let
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   461
    fun findParallelNeq []  = NONE
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   462
    |   findParallelNeq (e::es)  =
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   463
     if (x aconv (lower e) andalso y aconv (upper e)) then SOME e
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   464
     else if (y aconv (lower e) andalso x aconv (upper e))
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   465
     then SOME (NotEq (x,y, (Thm ([getprf e], @{thm not_sym}))))
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   466
     else findParallelNeq es;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   467
   in
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   468
   (* test if there is a edge x ~= y respectivly  y ~= x and     *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   469
   (* if it possible to find a path x <= y in leG, thus we can conclude x < y *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   470
    (case findParallelNeq neqE of (SOME e) =>
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   471
      let
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   472
       val (xyLeFound,xyLePath) = findPath x y leG
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   473
      in
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   474
       if xyLeFound then (
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   475
        let
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   476
         val Le_x_y = (transPath (tl xyLePath, hd xyLePath))
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   477
         val Less_x_y = mergeLess (e, Le_x_y)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   478
        in getprf Less_x_y end
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   479
       ) else raise Cannot
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   480
      end
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   481
    | _ => raise Cannot)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   482
   end )
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   483
 | (NotEq (x,y,_)) =>
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   484
  (* First check if a single premiss is sufficient *)
59584
wenzelm
parents: 59498
diff changeset
   485
  (case (find_first (fn fact => fact subsumes subgoal) neqE, subgoal) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   486
    (SOME (NotEq (x, y, p)), NotEq (x', y', _)) =>
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   487
      if  (x aconv x' andalso y aconv y') then p
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 37744
diff changeset
   488
      else Thm ([p], @{thm not_sym})
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   489
    | _  => raise Cannot
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   490
  )
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   491
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   492
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   493
(* ************************************************************************ *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   494
(*                                                                          *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   495
(* End: Quasi Order relevant functions                                      *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   496
(*                                                                          *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   497
(*                                                                          *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   498
(* ************************************************************************ *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   499
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   500
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   501
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   502
(* solveLeTrans sign (asms,concl) :                                        *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   503
(* theory -> less list * Term.term -> proof list                           *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   504
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   505
(* Solves                                                                  *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   506
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   507
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   508
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   509
fun solveLeTrans thy (asms, concl) =
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   510
 let
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   511
  val g = mkGraph asms
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   512
 in
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   513
   let
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   514
    val (subgoal, prf) = mkconcl_trans thy concl
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   515
    val (found, path) = findPath (lower subgoal) (upper subgoal) g
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   516
   in
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   517
    if found then [getprf (transPath (tl path, hd path))]
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   518
    else raise Cannot
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   519
  end
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   520
 end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   521
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   522
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   523
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   524
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   525
(* solveQuasiOrder sign (asms,concl) :                                     *)
19250
932a50e2332f got rid of type Sign.sg;
wenzelm
parents: 15570
diff changeset
   526
(* theory -> less list * Term.term -> proof list                           *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   527
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   528
(* Find proof if possible for quasi order.                                 *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   529
(*                                                                         *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   530
(* *********************************************************************** *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   531
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   532
fun solveQuasiOrder thy (asms, concl) =
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   533
 let
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   534
  val (leG, neqE) = mkQuasiGraph asms
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   535
 in
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   536
   let
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   537
   val (subgoals, prf) = mkconcl_quasi thy concl
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   538
   fun solve facts less =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   539
       (case triv_solv less of NONE => findQuasiProof (leG, neqE) less
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15103
diff changeset
   540
       | SOME prf => prf )
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   541
  in   map (solve asms) subgoals end
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   542
 end;
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   543
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   544
(* ************************************************************************ *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   545
(*                                                                          *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   546
(* Tactics                                                                  *)
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   547
(*                                                                          *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   548
(*  - trans_tac                                                          *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   549
(*  - quasi_tac, solves quasi orders                                        *)
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   550
(* ************************************************************************ *)
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   551
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   552
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   553
(* trans_tac - solves transitivity chains over <= *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   554
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
   555
fun trans_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   556
 let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 39159
diff changeset
   557
  val thy = Proof_Context.theory_of ctxt;
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   558
  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   559
  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A);
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   560
  val C = subst_bounds (rfrees, Logic.strip_assums_concl A);
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   561
  val lesss = flat (map_index (mkasm_trans thy o swap) Hs);
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   562
  val prfs = solveLeTrans thy (lesss, C);
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   563
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   564
  val (subgoal, prf) = mkconcl_trans thy C;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   565
 in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   566
  Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   567
    let val thms = map (prove prems) prfs
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   568
    in resolve_tac ctxt' [prove thms prf] 1 end) ctxt n st
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   569
 end
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   570
 handle Contr p =>
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   571
    Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   572
      resolve_tac ctxt' [prove prems p] 1) ctxt n st
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   573
  | Cannot => Seq.empty);
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   574
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   575
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   576
(* quasi_tac - solves quasi orders *)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   577
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   578
fun quasi_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   579
 let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 39159
diff changeset
   580
  val thy = Proof_Context.theory_of ctxt
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   581
  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   582
  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A);
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   583
  val C = subst_bounds (rfrees, Logic.strip_assums_concl A);
33063
4d462963a7db map_range (and map_index) combinator
haftmann
parents: 32952
diff changeset
   584
  val lesss = flat (map_index (mkasm_quasi thy o swap) Hs);
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   585
  val prfs = solveQuasiOrder thy (lesss, C);
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   586
  val (subgoals, prf) = mkconcl_quasi thy C;
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   587
 in
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   588
  Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   589
    let val thms = map (prove prems) prfs
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   590
    in resolve_tac ctxt' [prove thms prf] 1 end) ctxt n st
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   591
 end
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   592
 handle Contr p =>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   593
    (Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58839
diff changeset
   594
      resolve_tac ctxt' [prove prems p] 1) ctxt n st handle General.Subscript => Seq.empty)
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   595
  | Cannot => Seq.empty
43278
1fbdcebb364b more robust exception pattern General.Subscript;
wenzelm
parents: 42364
diff changeset
   596
  | General.Subscript => Seq.empty);
32215
87806301a813 replaced old METAHYPS by FOCUS;
wenzelm
parents: 29276
diff changeset
   597
15103
79846e8792eb New transitivity reasoners for transitivity only and quasi orders.
ballarin
parents:
diff changeset
   598
end;