src/HOL/Library/Old_SMT/old_z3_proof_literals.ML
changeset 65515 f595b7532dc9
parent 65514 d10f0bbc7ea1
child 65516 03efd17e083b
equal deleted inserted replaced
65514:d10f0bbc7ea1 65515:f595b7532dc9
     1 (*  Title:      HOL/Library/Old_SMT/old_z3_proof_literals.ML
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 
       
     4 Proof tools related to conjunctions and disjunctions.
       
     5 *)
       
     6 
       
     7 signature OLD_Z3_PROOF_LITERALS =
       
     8 sig
       
     9   (*literal table*)
       
    10   type littab = thm Termtab.table
       
    11   val make_littab: thm list -> littab
       
    12   val insert_lit: thm -> littab -> littab
       
    13   val delete_lit: thm -> littab -> littab
       
    14   val lookup_lit: littab -> term -> thm option
       
    15   val get_first_lit: (term -> bool) -> littab -> thm option
       
    16 
       
    17   (*rules*)
       
    18   val true_thm: thm
       
    19   val rewrite_true: thm
       
    20 
       
    21   (*properties*)
       
    22   val is_conj: term -> bool
       
    23   val is_disj: term -> bool
       
    24   val exists_lit: bool -> (term -> bool) -> term -> bool
       
    25   val negate: cterm -> cterm
       
    26 
       
    27   (*proof tools*)
       
    28   val explode: bool -> bool -> bool -> term list -> thm -> thm list
       
    29   val join: bool -> littab -> term -> thm
       
    30   val prove_conj_disj_eq: cterm -> thm
       
    31 end
       
    32 
       
    33 structure Old_Z3_Proof_Literals: OLD_Z3_PROOF_LITERALS =
       
    34 struct
       
    35 
       
    36 
       
    37 
       
    38 (* literal table *)
       
    39 
       
    40 type littab = thm Termtab.table
       
    41 
       
    42 fun make_littab thms =
       
    43   fold (Termtab.update o `Old_SMT_Utils.prop_of) thms Termtab.empty
       
    44 
       
    45 fun insert_lit thm = Termtab.update (`Old_SMT_Utils.prop_of thm)
       
    46 fun delete_lit thm = Termtab.delete (Old_SMT_Utils.prop_of thm)
       
    47 fun lookup_lit lits = Termtab.lookup lits
       
    48 fun get_first_lit f =
       
    49   Termtab.get_first (fn (t, thm) => if f t then SOME thm else NONE)
       
    50 
       
    51 
       
    52 
       
    53 (* rules *)
       
    54 
       
    55 val true_thm = @{lemma "~False" by simp}
       
    56 val rewrite_true = @{lemma "True == ~ False" by simp}
       
    57 
       
    58 
       
    59 
       
    60 (* properties and term operations *)
       
    61 
       
    62 val is_neg = (fn @{const Not} $ _ => true | _ => false)
       
    63 fun is_neg' f = (fn @{const Not} $ t => f t | _ => false)
       
    64 val is_dneg = is_neg' is_neg
       
    65 val is_conj = (fn @{const HOL.conj} $ _ $ _ => true | _ => false)
       
    66 val is_disj = (fn @{const HOL.disj} $ _ $ _ => true | _ => false)
       
    67 
       
    68 fun dest_disj_term' f = (fn
       
    69     @{const Not} $ (@{const HOL.disj} $ t $ u) => SOME (f t, f u)
       
    70   | _ => NONE)
       
    71 
       
    72 val dest_conj_term = (fn @{const HOL.conj} $ t $ u => SOME (t, u) | _ => NONE)
       
    73 val dest_disj_term =
       
    74   dest_disj_term' (fn @{const Not} $ t => t | t => @{const Not} $ t)
       
    75 
       
    76 fun exists_lit is_conj P =
       
    77   let
       
    78     val dest = if is_conj then dest_conj_term else dest_disj_term
       
    79     fun exists t = P t orelse
       
    80       (case dest t of
       
    81         SOME (t1, t2) => exists t1 orelse exists t2
       
    82       | NONE => false)
       
    83   in exists end
       
    84 
       
    85 val negate = Thm.apply (Thm.cterm_of @{context} @{const Not})
       
    86 
       
    87 
       
    88 
       
    89 (* proof tools *)
       
    90 
       
    91 (** explosion of conjunctions and disjunctions **)
       
    92 
       
    93 local
       
    94   val precomp = Old_Z3_Proof_Tools.precompose2
       
    95 
       
    96   fun destc ct = Thm.dest_binop (Thm.dest_arg ct)
       
    97   val dest_conj1 = precomp destc @{thm conjunct1}
       
    98   val dest_conj2 = precomp destc @{thm conjunct2}
       
    99   fun dest_conj_rules t =
       
   100     dest_conj_term t |> Option.map (K (dest_conj1, dest_conj2))
       
   101     
       
   102   fun destd f ct = f (Thm.dest_binop (Thm.dest_arg (Thm.dest_arg ct)))
       
   103   val dn1 = apfst Thm.dest_arg and dn2 = apsnd Thm.dest_arg
       
   104   val dest_disj1 = precomp (destd I) @{lemma "~(P | Q) ==> ~P" by fast}
       
   105   val dest_disj2 = precomp (destd dn1) @{lemma "~(~P | Q) ==> P" by fast}
       
   106   val dest_disj3 = precomp (destd I) @{lemma "~(P | Q) ==> ~Q" by fast}
       
   107   val dest_disj4 = precomp (destd dn2) @{lemma "~(P | ~Q) ==> Q" by fast}
       
   108 
       
   109   fun dest_disj_rules t =
       
   110     (case dest_disj_term' is_neg t of
       
   111       SOME (true, true) => SOME (dest_disj2, dest_disj4)
       
   112     | SOME (true, false) => SOME (dest_disj2, dest_disj3)
       
   113     | SOME (false, true) => SOME (dest_disj1, dest_disj4)
       
   114     | SOME (false, false) => SOME (dest_disj1, dest_disj3)
       
   115     | NONE => NONE)
       
   116 
       
   117   fun destn ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg ct))]
       
   118   val dneg_rule = Old_Z3_Proof_Tools.precompose destn @{thm notnotD}
       
   119 in
       
   120 
       
   121 (*
       
   122   explode a term into literals and collect all rules to be able to deduce
       
   123   particular literals afterwards
       
   124 *)
       
   125 fun explode_term is_conj =
       
   126   let
       
   127     val dest = if is_conj then dest_conj_term else dest_disj_term
       
   128     val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules
       
   129 
       
   130     fun add (t, rs) = Termtab.map_default (t, rs)
       
   131       (fn rs' => if length rs' < length rs then rs' else rs)
       
   132 
       
   133     fun explode1 rules t =
       
   134       (case dest t of
       
   135         SOME (t1, t2) =>
       
   136           let val (rule1, rule2) = the (dest_rules t)
       
   137           in
       
   138             explode1 (rule1 :: rules) t1 #>
       
   139             explode1 (rule2 :: rules) t2 #>
       
   140             add (t, rev rules)
       
   141           end
       
   142       | NONE => add (t, rev rules))
       
   143 
       
   144     fun explode0 (@{const Not} $ (@{const Not} $ t)) =
       
   145           Termtab.make [(t, [dneg_rule])]
       
   146       | explode0 t = explode1 [] t Termtab.empty
       
   147 
       
   148   in explode0 end
       
   149 
       
   150 (*
       
   151   extract a literal by applying previously collected rules
       
   152 *)
       
   153 fun extract_lit thm rules = fold Old_Z3_Proof_Tools.compose rules thm
       
   154 
       
   155 
       
   156 (*
       
   157   explode a theorem into its literals
       
   158 *)
       
   159 fun explode is_conj full keep_intermediate stop_lits =
       
   160   let
       
   161     val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules
       
   162     val tab = fold (Termtab.update o rpair ()) stop_lits Termtab.empty
       
   163 
       
   164     fun explode1 thm =
       
   165       if Termtab.defined tab (Old_SMT_Utils.prop_of thm) then cons thm
       
   166       else
       
   167         (case dest_rules (Old_SMT_Utils.prop_of thm) of
       
   168           SOME (rule1, rule2) =>
       
   169             explode2 rule1 thm #>
       
   170             explode2 rule2 thm #>
       
   171             keep_intermediate ? cons thm
       
   172         | NONE => cons thm)
       
   173 
       
   174     and explode2 dest_rule thm =
       
   175       if full orelse
       
   176         exists_lit is_conj (Termtab.defined tab) (Old_SMT_Utils.prop_of thm)
       
   177       then explode1 (Old_Z3_Proof_Tools.compose dest_rule thm)
       
   178       else cons (Old_Z3_Proof_Tools.compose dest_rule thm)
       
   179 
       
   180     fun explode0 thm =
       
   181       if not is_conj andalso is_dneg (Old_SMT_Utils.prop_of thm)
       
   182       then [Old_Z3_Proof_Tools.compose dneg_rule thm]
       
   183       else explode1 thm []
       
   184 
       
   185   in explode0 end
       
   186 
       
   187 end
       
   188 
       
   189 
       
   190 
       
   191 (** joining of literals to conjunctions or disjunctions **)
       
   192 
       
   193 local
       
   194   fun on_cprem i f thm = f (Thm.cprem_of thm i)
       
   195   fun on_cprop f thm = f (Thm.cprop_of thm)
       
   196   fun precomp2 f g thm = (on_cprem 1 f thm, on_cprem 2 g thm, f, g, thm)
       
   197   fun comp2 (cv1, cv2, f, g, rule) thm1 thm2 =
       
   198     Thm.instantiate ([],
       
   199       [(dest_Var (Thm.term_of cv1), on_cprop f thm1),
       
   200        (dest_Var (Thm.term_of cv2), on_cprop g thm2)]) rule
       
   201     |> Old_Z3_Proof_Tools.discharge thm1 |> Old_Z3_Proof_Tools.discharge thm2
       
   202 
       
   203   fun d1 ct = Thm.dest_arg ct and d2 ct = Thm.dest_arg (Thm.dest_arg ct)
       
   204 
       
   205   val conj_rule = precomp2 d1 d1 @{thm conjI}
       
   206   fun comp_conj ((_, thm1), (_, thm2)) = comp2 conj_rule thm1 thm2
       
   207 
       
   208   val disj1 = precomp2 d2 d2 @{lemma "~P ==> ~Q ==> ~(P | Q)" by fast}
       
   209   val disj2 = precomp2 d2 d1 @{lemma "~P ==> Q ==> ~(P | ~Q)" by fast}
       
   210   val disj3 = precomp2 d1 d2 @{lemma "P ==> ~Q ==> ~(~P | Q)" by fast}
       
   211   val disj4 = precomp2 d1 d1 @{lemma "P ==> Q ==> ~(~P | ~Q)" by fast}
       
   212 
       
   213   fun comp_disj ((false, thm1), (false, thm2)) = comp2 disj1 thm1 thm2
       
   214     | comp_disj ((false, thm1), (true, thm2)) = comp2 disj2 thm1 thm2
       
   215     | comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2
       
   216     | comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2
       
   217 
       
   218   fun dest_conj (@{const HOL.conj} $ t $ u) = ((false, t), (false, u))
       
   219     | dest_conj t = raise TERM ("dest_conj", [t])
       
   220 
       
   221   val neg = (fn @{const Not} $ t => (true, t) | t => (false, @{const Not} $ t))
       
   222   fun dest_disj (@{const Not} $ (@{const HOL.disj} $ t $ u)) = (neg t, neg u)
       
   223     | dest_disj t = raise TERM ("dest_disj", [t])
       
   224 
       
   225   val precomp = Old_Z3_Proof_Tools.precompose
       
   226   val dnegE = precomp (single o d2 o d1) @{thm notnotD}
       
   227   val dnegI = precomp (single o d1) @{lemma "P ==> ~~P" by fast}
       
   228   fun as_dneg f t = f (@{const Not} $ (@{const Not} $ t))
       
   229 
       
   230   val precomp2 = Old_Z3_Proof_Tools.precompose2
       
   231   fun dni f = apsnd f o Thm.dest_binop o f o d1
       
   232   val negIffE = precomp2 (dni d1) @{lemma "~(P = (~Q)) ==> Q = P" by fast}
       
   233   val negIffI = precomp2 (dni I) @{lemma "P = Q ==> ~(Q = (~P))" by fast}
       
   234   val iff_const = @{const HOL.eq (bool)}
       
   235   fun as_negIff f (@{const HOL.eq (bool)} $ t $ u) =
       
   236         f (@{const Not} $ (iff_const $ u $ (@{const Not} $ t)))
       
   237     | as_negIff _ _ = NONE
       
   238 in
       
   239 
       
   240 fun join is_conj littab t =
       
   241   let
       
   242     val comp = if is_conj then comp_conj else comp_disj
       
   243     val dest = if is_conj then dest_conj else dest_disj
       
   244 
       
   245     val lookup = lookup_lit littab
       
   246 
       
   247     fun lookup_rule t =
       
   248       (case t of
       
   249         @{const Not} $ (@{const Not} $ t) =>
       
   250           (Old_Z3_Proof_Tools.compose dnegI, lookup t)
       
   251       | @{const Not} $ (@{const HOL.eq (bool)} $ t $ (@{const Not} $ u)) =>
       
   252           (Old_Z3_Proof_Tools.compose negIffI, lookup (iff_const $ u $ t))
       
   253       | @{const Not} $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) =>
       
   254           let fun rewr lit = lit COMP @{thm not_sym}
       
   255           in (rewr, lookup (@{const Not} $ (eq $ u $ t))) end
       
   256       | _ =>
       
   257           (case as_dneg lookup t of
       
   258             NONE => (Old_Z3_Proof_Tools.compose negIffE, as_negIff lookup t)
       
   259           | x => (Old_Z3_Proof_Tools.compose dnegE, x)))
       
   260 
       
   261     fun join1 (s, t) =
       
   262       (case lookup t of
       
   263         SOME lit => (s, lit)
       
   264       | NONE => 
       
   265           (case lookup_rule t of
       
   266             (rewrite, SOME lit) => (s, rewrite lit)
       
   267           | (_, NONE) => (s, comp (apply2 join1 (dest t)))))
       
   268 
       
   269   in snd (join1 (if is_conj then (false, t) else (true, t))) end
       
   270 
       
   271 end
       
   272 
       
   273 
       
   274 
       
   275 (** proving equality of conjunctions or disjunctions **)
       
   276 
       
   277 fun iff_intro thm1 thm2 = thm2 COMP (thm1 COMP @{thm iffI})
       
   278 
       
   279 local
       
   280   val cp1 = @{lemma "(~P) = (~Q) ==> P = Q" by simp}
       
   281   val cp2 = @{lemma "(~P) = Q ==> P = (~Q)" by fastforce}
       
   282   val cp3 = @{lemma "P = (~Q) ==> (~P) = Q" by simp}
       
   283 in
       
   284 fun contrapos1 prove (ct, cu) = prove (negate ct, negate cu) COMP cp1
       
   285 fun contrapos2 prove (ct, cu) = prove (negate ct, Thm.dest_arg cu) COMP cp2
       
   286 fun contrapos3 prove (ct, cu) = prove (Thm.dest_arg ct, negate cu) COMP cp3
       
   287 end
       
   288 
       
   289 
       
   290 local
       
   291   val contra_rule = @{lemma "P ==> ~P ==> False" by (rule notE)}
       
   292   fun contra_left conj thm =
       
   293     let
       
   294       val rules = explode_term conj (Old_SMT_Utils.prop_of thm)
       
   295       fun contra_lits (t, rs) =
       
   296         (case t of
       
   297           @{const Not} $ u => Termtab.lookup rules u |> Option.map (pair rs)
       
   298         | _ => NONE)
       
   299     in
       
   300       (case Termtab.lookup rules @{const False} of
       
   301         SOME rs => extract_lit thm rs
       
   302       | NONE =>
       
   303           the (Termtab.get_first contra_lits rules)
       
   304           |> apply2 (extract_lit thm)
       
   305           |> (fn (nlit, plit) => nlit COMP (plit COMP contra_rule)))
       
   306     end
       
   307 
       
   308   val falseE_v = dest_Var (Thm.term_of (Thm.dest_arg (Thm.dest_arg (Thm.cprop_of @{thm FalseE}))))
       
   309   fun contra_right ct = Thm.instantiate ([], [(falseE_v, ct)]) @{thm FalseE}
       
   310 in
       
   311 fun contradict conj ct =
       
   312   iff_intro (Old_Z3_Proof_Tools.under_assumption (contra_left conj) ct)
       
   313     (contra_right ct)
       
   314 end
       
   315 
       
   316 
       
   317 local
       
   318   fun prove_eq l r (cl, cr) =
       
   319     let
       
   320       fun explode' is_conj = explode is_conj true (l <> r) []
       
   321       fun make_tab is_conj thm = make_littab (true_thm :: explode' is_conj thm)
       
   322       fun prove is_conj ct tab = join is_conj tab (Thm.term_of ct)
       
   323 
       
   324       val thm1 = Old_Z3_Proof_Tools.under_assumption (prove r cr o make_tab l) cl
       
   325       val thm2 = Old_Z3_Proof_Tools.under_assumption (prove l cl o make_tab r) cr
       
   326     in iff_intro thm1 thm2 end
       
   327 
       
   328   datatype conj_disj = CONJ | DISJ | NCON | NDIS
       
   329   fun kind_of t =
       
   330     if is_conj t then SOME CONJ
       
   331     else if is_disj t then SOME DISJ
       
   332     else if is_neg' is_conj t then SOME NCON
       
   333     else if is_neg' is_disj t then SOME NDIS
       
   334     else NONE
       
   335 in
       
   336 
       
   337 fun prove_conj_disj_eq ct =
       
   338   let val cp as (cl, cr) = Thm.dest_binop (Thm.dest_arg ct)
       
   339   in
       
   340     (case (kind_of (Thm.term_of cl), Thm.term_of cr) of
       
   341       (SOME CONJ, @{const False}) => contradict true cl
       
   342     | (SOME DISJ, @{const Not} $ @{const False}) =>
       
   343         contrapos2 (contradict false o fst) cp
       
   344     | (kl, _) =>
       
   345         (case (kl, kind_of (Thm.term_of cr)) of
       
   346           (SOME CONJ, SOME CONJ) => prove_eq true true cp
       
   347         | (SOME CONJ, SOME NDIS) => prove_eq true false cp
       
   348         | (SOME CONJ, _) => prove_eq true true cp
       
   349         | (SOME DISJ, SOME DISJ) => contrapos1 (prove_eq false false) cp
       
   350         | (SOME DISJ, SOME NCON) => contrapos2 (prove_eq false true) cp
       
   351         | (SOME DISJ, _) => contrapos1 (prove_eq false false) cp
       
   352         | (SOME NCON, SOME NCON) => contrapos1 (prove_eq true true) cp
       
   353         | (SOME NCON, SOME DISJ) => contrapos3 (prove_eq true false) cp
       
   354         | (SOME NCON, NONE) => contrapos3 (prove_eq true false) cp
       
   355         | (SOME NDIS, SOME NDIS) => prove_eq false false cp
       
   356         | (SOME NDIS, SOME CONJ) => prove_eq false true cp
       
   357         | (SOME NDIS, NONE) => prove_eq false true cp
       
   358         | _ => raise CTERM ("prove_conj_disj_eq", [ct])))
       
   359   end
       
   360 
       
   361 end
       
   362 
       
   363 end