blanchet@58058: (* Title: HOL/Library/Old_SMT/old_z3_proof_literals.ML boehmes@36898: Author: Sascha Boehme, TU Muenchen boehmes@36898: boehmes@36898: Proof tools related to conjunctions and disjunctions. boehmes@36898: *) boehmes@36898: blanchet@58058: signature OLD_Z3_PROOF_LITERALS = boehmes@36898: sig boehmes@41123: (*literal table*) boehmes@36898: type littab = thm Termtab.table boehmes@36898: val make_littab: thm list -> littab boehmes@36898: val insert_lit: thm -> littab -> littab boehmes@36898: val delete_lit: thm -> littab -> littab boehmes@36898: val lookup_lit: littab -> term -> thm option boehmes@36898: val get_first_lit: (term -> bool) -> littab -> thm option boehmes@36898: boehmes@41123: (*rules*) boehmes@36898: val true_thm: thm boehmes@36898: val rewrite_true: thm boehmes@36898: boehmes@41123: (*properties*) boehmes@36898: val is_conj: term -> bool boehmes@36898: val is_disj: term -> bool boehmes@36898: val exists_lit: bool -> (term -> bool) -> term -> bool boehmes@40579: val negate: cterm -> cterm boehmes@36898: boehmes@41123: (*proof tools*) boehmes@36898: val explode: bool -> bool -> bool -> term list -> thm -> thm list boehmes@36898: val join: bool -> littab -> term -> thm boehmes@36898: val prove_conj_disj_eq: cterm -> thm boehmes@36898: end boehmes@36898: blanchet@58058: structure Old_Z3_Proof_Literals: OLD_Z3_PROOF_LITERALS = boehmes@36898: struct boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41123: (* literal table *) boehmes@36898: boehmes@36898: type littab = thm Termtab.table boehmes@36898: boehmes@41328: fun make_littab thms = blanchet@58058: fold (Termtab.update o `Old_SMT_Utils.prop_of) thms Termtab.empty boehmes@36898: blanchet@58058: fun insert_lit thm = Termtab.update (`Old_SMT_Utils.prop_of thm) blanchet@58058: fun delete_lit thm = Termtab.delete (Old_SMT_Utils.prop_of thm) boehmes@36898: fun lookup_lit lits = Termtab.lookup lits boehmes@36898: fun get_first_lit f = boehmes@36898: Termtab.get_first (fn (t, thm) => if f t then SOME thm else NONE) boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41123: (* rules *) boehmes@36898: boehmes@36898: val true_thm = @{lemma "~False" by simp} boehmes@36898: val rewrite_true = @{lemma "True == ~ False" by simp} boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41123: (* properties and term operations *) boehmes@36898: boehmes@40579: val is_neg = (fn @{const Not} $ _ => true | _ => false) boehmes@40579: fun is_neg' f = (fn @{const Not} $ t => f t | _ => false) boehmes@36898: val is_dneg = is_neg' is_neg boehmes@40579: val is_conj = (fn @{const HOL.conj} $ _ $ _ => true | _ => false) boehmes@40579: val is_disj = (fn @{const HOL.disj} $ _ $ _ => true | _ => false) boehmes@36898: boehmes@36898: fun dest_disj_term' f = (fn boehmes@40579: @{const Not} $ (@{const HOL.disj} $ t $ u) => SOME (f t, f u) boehmes@36898: | _ => NONE) boehmes@36898: boehmes@40579: val dest_conj_term = (fn @{const HOL.conj} $ t $ u => SOME (t, u) | _ => NONE) boehmes@36898: val dest_disj_term = boehmes@40579: dest_disj_term' (fn @{const Not} $ t => t | t => @{const Not} $ t) boehmes@36898: boehmes@36898: fun exists_lit is_conj P = boehmes@36898: let boehmes@36898: val dest = if is_conj then dest_conj_term else dest_disj_term boehmes@36898: fun exists t = P t orelse boehmes@36898: (case dest t of boehmes@36898: SOME (t1, t2) => exists t1 orelse exists t2 boehmes@36898: | NONE => false) boehmes@36898: in exists end boehmes@36898: wenzelm@46497: val negate = Thm.apply (Thm.cterm_of @{theory} @{const Not}) boehmes@40579: boehmes@36898: boehmes@36898: boehmes@41123: (* proof tools *) boehmes@36898: boehmes@41123: (** explosion of conjunctions and disjunctions **) boehmes@36898: boehmes@36898: local blanchet@58058: val precomp = Old_Z3_Proof_Tools.precompose2 boehmes@41328: boehmes@36898: fun destc ct = Thm.dest_binop (Thm.dest_arg ct) boehmes@41328: val dest_conj1 = precomp destc @{thm conjunct1} boehmes@41328: val dest_conj2 = precomp destc @{thm conjunct2} boehmes@36898: fun dest_conj_rules t = boehmes@36898: dest_conj_term t |> Option.map (K (dest_conj1, dest_conj2)) boehmes@36898: boehmes@36898: fun destd f ct = f (Thm.dest_binop (Thm.dest_arg (Thm.dest_arg ct))) boehmes@36898: val dn1 = apfst Thm.dest_arg and dn2 = apsnd Thm.dest_arg boehmes@41328: val dest_disj1 = precomp (destd I) @{lemma "~(P | Q) ==> ~P" by fast} boehmes@41328: val dest_disj2 = precomp (destd dn1) @{lemma "~(~P | Q) ==> P" by fast} boehmes@41328: val dest_disj3 = precomp (destd I) @{lemma "~(P | Q) ==> ~Q" by fast} boehmes@41328: val dest_disj4 = precomp (destd dn2) @{lemma "~(P | ~Q) ==> Q" by fast} boehmes@36898: boehmes@36898: fun dest_disj_rules t = boehmes@36898: (case dest_disj_term' is_neg t of boehmes@36898: SOME (true, true) => SOME (dest_disj2, dest_disj4) boehmes@36898: | SOME (true, false) => SOME (dest_disj2, dest_disj3) boehmes@36898: | SOME (false, true) => SOME (dest_disj1, dest_disj4) boehmes@36898: | SOME (false, false) => SOME (dest_disj1, dest_disj3) boehmes@36898: | NONE => NONE) boehmes@36898: boehmes@36898: fun destn ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg ct))] blanchet@58058: val dneg_rule = Old_Z3_Proof_Tools.precompose destn @{thm notnotD} boehmes@36898: in boehmes@36898: boehmes@41123: (* boehmes@41123: explode a term into literals and collect all rules to be able to deduce boehmes@41123: particular literals afterwards boehmes@41123: *) boehmes@36898: fun explode_term is_conj = boehmes@36898: let boehmes@36898: val dest = if is_conj then dest_conj_term else dest_disj_term boehmes@36898: val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules boehmes@36898: boehmes@36898: fun add (t, rs) = Termtab.map_default (t, rs) boehmes@36898: (fn rs' => if length rs' < length rs then rs' else rs) boehmes@36898: boehmes@36898: fun explode1 rules t = boehmes@36898: (case dest t of boehmes@36898: SOME (t1, t2) => boehmes@36898: let val (rule1, rule2) = the (dest_rules t) boehmes@36898: in boehmes@36898: explode1 (rule1 :: rules) t1 #> boehmes@36898: explode1 (rule2 :: rules) t2 #> boehmes@36898: add (t, rev rules) boehmes@36898: end boehmes@36898: | NONE => add (t, rev rules)) boehmes@36898: boehmes@40579: fun explode0 (@{const Not} $ (@{const Not} $ t)) = boehmes@36898: Termtab.make [(t, [dneg_rule])] boehmes@36898: | explode0 t = explode1 [] t Termtab.empty boehmes@36898: boehmes@36898: in explode0 end boehmes@36898: boehmes@41123: (* boehmes@41123: extract a literal by applying previously collected rules boehmes@41123: *) blanchet@58058: fun extract_lit thm rules = fold Old_Z3_Proof_Tools.compose rules thm boehmes@36898: boehmes@36898: boehmes@41123: (* boehmes@41123: explode a theorem into its literals boehmes@41123: *) boehmes@36898: fun explode is_conj full keep_intermediate stop_lits = boehmes@36898: let boehmes@36898: val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules boehmes@36898: val tab = fold (Termtab.update o rpair ()) stop_lits Termtab.empty boehmes@36898: boehmes@36898: fun explode1 thm = blanchet@58058: if Termtab.defined tab (Old_SMT_Utils.prop_of thm) then cons thm boehmes@36898: else blanchet@58058: (case dest_rules (Old_SMT_Utils.prop_of thm) of boehmes@36898: SOME (rule1, rule2) => boehmes@36898: explode2 rule1 thm #> boehmes@36898: explode2 rule2 thm #> boehmes@36898: keep_intermediate ? cons thm boehmes@36898: | NONE => cons thm) boehmes@36898: boehmes@36898: and explode2 dest_rule thm = boehmes@41328: if full orelse blanchet@58058: exists_lit is_conj (Termtab.defined tab) (Old_SMT_Utils.prop_of thm) blanchet@58058: then explode1 (Old_Z3_Proof_Tools.compose dest_rule thm) blanchet@58058: else cons (Old_Z3_Proof_Tools.compose dest_rule thm) boehmes@36898: boehmes@36898: fun explode0 thm = blanchet@58058: if not is_conj andalso is_dneg (Old_SMT_Utils.prop_of thm) blanchet@58058: then [Old_Z3_Proof_Tools.compose dneg_rule thm] boehmes@36898: else explode1 thm [] boehmes@36898: boehmes@36898: in explode0 end boehmes@36898: boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41123: (** joining of literals to conjunctions or disjunctions **) boehmes@36898: boehmes@36898: local boehmes@36898: fun on_cprem i f thm = f (Thm.cprem_of thm i) boehmes@36898: fun on_cprop f thm = f (Thm.cprop_of thm) boehmes@36898: fun precomp2 f g thm = (on_cprem 1 f thm, on_cprem 2 g thm, f, g, thm) boehmes@36898: fun comp2 (cv1, cv2, f, g, rule) thm1 thm2 = boehmes@36898: Thm.instantiate ([], [(cv1, on_cprop f thm1), (cv2, on_cprop g thm2)]) rule blanchet@58058: |> Old_Z3_Proof_Tools.discharge thm1 |> Old_Z3_Proof_Tools.discharge thm2 boehmes@36898: boehmes@36898: fun d1 ct = Thm.dest_arg ct and d2 ct = Thm.dest_arg (Thm.dest_arg ct) boehmes@36898: boehmes@36898: val conj_rule = precomp2 d1 d1 @{thm conjI} boehmes@36898: fun comp_conj ((_, thm1), (_, thm2)) = comp2 conj_rule thm1 thm2 boehmes@36898: boehmes@36898: val disj1 = precomp2 d2 d2 @{lemma "~P ==> ~Q ==> ~(P | Q)" by fast} boehmes@36898: val disj2 = precomp2 d2 d1 @{lemma "~P ==> Q ==> ~(P | ~Q)" by fast} boehmes@36898: val disj3 = precomp2 d1 d2 @{lemma "P ==> ~Q ==> ~(~P | Q)" by fast} boehmes@36898: val disj4 = precomp2 d1 d1 @{lemma "P ==> Q ==> ~(~P | ~Q)" by fast} boehmes@36898: boehmes@36898: fun comp_disj ((false, thm1), (false, thm2)) = comp2 disj1 thm1 thm2 boehmes@36898: | comp_disj ((false, thm1), (true, thm2)) = comp2 disj2 thm1 thm2 boehmes@36898: | comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2 boehmes@36898: | comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2 boehmes@36898: boehmes@40579: fun dest_conj (@{const HOL.conj} $ t $ u) = ((false, t), (false, u)) boehmes@36898: | dest_conj t = raise TERM ("dest_conj", [t]) boehmes@36898: boehmes@40579: val neg = (fn @{const Not} $ t => (true, t) | t => (false, @{const Not} $ t)) boehmes@40579: fun dest_disj (@{const Not} $ (@{const HOL.disj} $ t $ u)) = (neg t, neg u) boehmes@36898: | dest_disj t = raise TERM ("dest_disj", [t]) boehmes@36898: blanchet@58058: val precomp = Old_Z3_Proof_Tools.precompose boehmes@41328: val dnegE = precomp (single o d2 o d1) @{thm notnotD} boehmes@41328: val dnegI = precomp (single o d1) @{lemma "P ==> ~~P" by fast} boehmes@40579: fun as_dneg f t = f (@{const Not} $ (@{const Not} $ t)) boehmes@36898: blanchet@58058: val precomp2 = Old_Z3_Proof_Tools.precompose2 boehmes@36898: fun dni f = apsnd f o Thm.dest_binop o f o d1 boehmes@41328: val negIffE = precomp2 (dni d1) @{lemma "~(P = (~Q)) ==> Q = P" by fast} boehmes@41328: val negIffI = precomp2 (dni I) @{lemma "P = Q ==> ~(Q = (~P))" by fast} boehmes@40579: val iff_const = @{const HOL.eq (bool)} boehmes@40579: fun as_negIff f (@{const HOL.eq (bool)} $ t $ u) = boehmes@40579: f (@{const Not} $ (iff_const $ u $ (@{const Not} $ t))) boehmes@36898: | as_negIff _ _ = NONE boehmes@36898: in boehmes@36898: boehmes@36898: fun join is_conj littab t = boehmes@36898: let boehmes@36898: val comp = if is_conj then comp_conj else comp_disj boehmes@36898: val dest = if is_conj then dest_conj else dest_disj boehmes@36898: boehmes@36898: val lookup = lookup_lit littab boehmes@36898: boehmes@36898: fun lookup_rule t = boehmes@36898: (case t of boehmes@41328: @{const Not} $ (@{const Not} $ t) => blanchet@58058: (Old_Z3_Proof_Tools.compose dnegI, lookup t) boehmes@40579: | @{const Not} $ (@{const HOL.eq (bool)} $ t $ (@{const Not} $ u)) => blanchet@58058: (Old_Z3_Proof_Tools.compose negIffI, lookup (iff_const $ u $ t)) boehmes@40579: | @{const Not} $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) => boehmes@36898: let fun rewr lit = lit COMP @{thm not_sym} boehmes@40579: in (rewr, lookup (@{const Not} $ (eq $ u $ t))) end boehmes@36898: | _ => boehmes@36898: (case as_dneg lookup t of blanchet@58058: NONE => (Old_Z3_Proof_Tools.compose negIffE, as_negIff lookup t) blanchet@58058: | x => (Old_Z3_Proof_Tools.compose dnegE, x))) boehmes@36898: boehmes@36898: fun join1 (s, t) = boehmes@36898: (case lookup t of boehmes@36898: SOME lit => (s, lit) boehmes@36898: | NONE => boehmes@36898: (case lookup_rule t of boehmes@36898: (rewrite, SOME lit) => (s, rewrite lit) boehmes@36898: | (_, NONE) => (s, comp (pairself join1 (dest t))))) boehmes@36898: boehmes@36898: in snd (join1 (if is_conj then (false, t) else (true, t))) end boehmes@36898: boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41123: (** proving equality of conjunctions or disjunctions **) boehmes@36898: boehmes@36898: fun iff_intro thm1 thm2 = thm2 COMP (thm1 COMP @{thm iffI}) boehmes@36898: boehmes@36898: local boehmes@36898: val cp1 = @{lemma "(~P) = (~Q) ==> P = Q" by simp} nipkow@44890: val cp2 = @{lemma "(~P) = Q ==> P = (~Q)" by fastforce} boehmes@36898: val cp3 = @{lemma "P = (~Q) ==> (~P) = Q" by simp} boehmes@36898: in boehmes@40579: fun contrapos1 prove (ct, cu) = prove (negate ct, negate cu) COMP cp1 boehmes@40579: fun contrapos2 prove (ct, cu) = prove (negate ct, Thm.dest_arg cu) COMP cp2 boehmes@40579: fun contrapos3 prove (ct, cu) = prove (Thm.dest_arg ct, negate cu) COMP cp3 boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: local boehmes@36898: val contra_rule = @{lemma "P ==> ~P ==> False" by (rule notE)} boehmes@36898: fun contra_left conj thm = boehmes@36898: let blanchet@58058: val rules = explode_term conj (Old_SMT_Utils.prop_of thm) boehmes@36898: fun contra_lits (t, rs) = boehmes@36898: (case t of boehmes@40579: @{const Not} $ u => Termtab.lookup rules u |> Option.map (pair rs) boehmes@36898: | _ => NONE) boehmes@36898: in boehmes@40579: (case Termtab.lookup rules @{const False} of boehmes@36898: SOME rs => extract_lit thm rs boehmes@36898: | NONE => boehmes@36898: the (Termtab.get_first contra_lits rules) boehmes@36898: |> pairself (extract_lit thm) boehmes@36898: |> (fn (nlit, plit) => nlit COMP (plit COMP contra_rule))) boehmes@36898: end boehmes@36898: boehmes@36898: val falseE_v = Thm.dest_arg (Thm.dest_arg (Thm.cprop_of @{thm FalseE})) boehmes@36898: fun contra_right ct = Thm.instantiate ([], [(falseE_v, ct)]) @{thm FalseE} boehmes@36898: in boehmes@36898: fun contradict conj ct = blanchet@58058: iff_intro (Old_Z3_Proof_Tools.under_assumption (contra_left conj) ct) boehmes@41328: (contra_right ct) boehmes@36898: end boehmes@36898: boehmes@36898: boehmes@36898: local boehmes@36898: fun prove_eq l r (cl, cr) = boehmes@36898: let boehmes@36898: fun explode' is_conj = explode is_conj true (l <> r) [] boehmes@36898: fun make_tab is_conj thm = make_littab (true_thm :: explode' is_conj thm) boehmes@36898: fun prove is_conj ct tab = join is_conj tab (Thm.term_of ct) boehmes@36898: blanchet@58058: val thm1 = Old_Z3_Proof_Tools.under_assumption (prove r cr o make_tab l) cl blanchet@58058: val thm2 = Old_Z3_Proof_Tools.under_assumption (prove l cl o make_tab r) cr boehmes@36898: in iff_intro thm1 thm2 end boehmes@36898: boehmes@36898: datatype conj_disj = CONJ | DISJ | NCON | NDIS boehmes@36898: fun kind_of t = boehmes@36898: if is_conj t then SOME CONJ boehmes@36898: else if is_disj t then SOME DISJ boehmes@36898: else if is_neg' is_conj t then SOME NCON boehmes@36898: else if is_neg' is_disj t then SOME NDIS boehmes@36898: else NONE boehmes@36898: in boehmes@36898: boehmes@36898: fun prove_conj_disj_eq ct = boehmes@36898: let val cp as (cl, cr) = Thm.dest_binop (Thm.dest_arg ct) boehmes@36898: in boehmes@36898: (case (kind_of (Thm.term_of cl), Thm.term_of cr) of boehmes@40579: (SOME CONJ, @{const False}) => contradict true cl boehmes@40579: | (SOME DISJ, @{const Not} $ @{const False}) => boehmes@40579: contrapos2 (contradict false o fst) cp boehmes@36898: | (kl, _) => boehmes@36898: (case (kl, kind_of (Thm.term_of cr)) of boehmes@36898: (SOME CONJ, SOME CONJ) => prove_eq true true cp boehmes@36898: | (SOME CONJ, SOME NDIS) => prove_eq true false cp boehmes@36898: | (SOME CONJ, _) => prove_eq true true cp boehmes@36898: | (SOME DISJ, SOME DISJ) => contrapos1 (prove_eq false false) cp boehmes@36898: | (SOME DISJ, SOME NCON) => contrapos2 (prove_eq false true) cp boehmes@36898: | (SOME DISJ, _) => contrapos1 (prove_eq false false) cp boehmes@36898: | (SOME NCON, SOME NCON) => contrapos1 (prove_eq true true) cp boehmes@36898: | (SOME NCON, SOME DISJ) => contrapos3 (prove_eq true false) cp boehmes@36898: | (SOME NCON, NONE) => contrapos3 (prove_eq true false) cp boehmes@36898: | (SOME NDIS, SOME NDIS) => prove_eq false false cp boehmes@36898: | (SOME NDIS, SOME CONJ) => prove_eq false true cp boehmes@36898: | (SOME NDIS, NONE) => prove_eq false true cp boehmes@36898: | _ => raise CTERM ("prove_conj_disj_eq", [ct]))) boehmes@36898: end boehmes@36898: boehmes@36898: end boehmes@36898: boehmes@36898: end