--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/activities/tphols09/IDW/CU-Ex1.thy Wed Mar 30 17:27:34 2016 +0100
@@ -0,0 +1,461 @@
+theory Ex
+imports Main
+begin
+
+text {*
+ An example of an apply-proof and the
+ corresponding ML-code.
+*}
+
+lemma disj_swap:
+ shows "P \<or> Q \<Longrightarrow> Q \<or> P"
+apply(erule disjE)
+apply(rule disjI2)
+apply(assumption)
+apply(rule disjI1)
+apply(assumption)
+done
+
+ML {*
+let
+ val ctxt = @{context}
+ val goal = @{prop "P \<or> Q \<Longrightarrow> Q \<or> P"}
+ val facts = []
+ val schms = ["P", "Q"]
+in
+ Goal.prove ctxt schms facts goal
+ (fn _ =>
+ etac @{thm disjE} 1
+ THEN rtac @{thm disjI2} 1
+ THEN atac 1
+ THEN rtac @{thm disjI1} 1
+ THEN atac 1)
+end
+*}
+
+text {*
+ Tactics coded in ML can be applied / tried out in
+ apply-scripts using "tactic". "THEN" is a tactic
+ combinator that just strings tactics together.
+ Tactic combinators are also called tacticals.
+*}
+
+ML {*
+val foo_tac =
+ (etac @{thm disjE} 1
+ THEN rtac @{thm disjI2} 1
+ THEN atac 1
+ THEN rtac @{thm disjI1} 1
+ THEN atac 1)
+*}
+
+lemma
+ shows "P \<or> Q \<Longrightarrow> Q \<or> P"
+apply(tactic {* foo_tac *})
+done
+
+text {*
+ Coding tactics with explicit subgoal addressing
+ is quite brittle and limits the usage of the tactics.
+ This can be avoided using the `primed' versions
+ of the tactic comibinators.
+*}
+
+ML {*
+val foo_tac' =
+ (etac @{thm disjE}
+ THEN' rtac @{thm disjI2}
+ THEN' atac
+ THEN' rtac @{thm disjI1}
+ THEN' atac)
+*}
+
+text {*
+ Now the tactic can be used on any subgoal.
+*}
+
+lemma "P1 \<or> Q1 \<Longrightarrow> Q1 \<or> P1"
+ and "P2 \<or> Q2 \<Longrightarrow> Q2 \<or> P2"
+apply(tactic {* foo_tac' 2 *})
+apply(tactic {* foo_tac' 1 *})
+done
+
+text {*
+ The type of a tactic is from a theorem
+ to a lazy sequence of (successor) theorems.
+
+ type tactic = thm -> thm Seq.seq
+
+ The simplest tactics are no_tac (always
+ fails) and all_tac (always succeds, but
+ does not make any progress.
+*}
+
+
+ML {* fun no_tac thm = Seq.empty *}
+ML {* fun all_tac thm = Seq.single thm *}
+
+text {*
+ The lazy sequence can be explored using
+ "back".
+*}
+
+lemma "\<lbrakk>P \<or> Q; P \<or> Q\<rbrakk> \<Longrightarrow> Q \<or> P"
+apply(tactic {* foo_tac' 1 *})
+back
+done
+
+text {*
+ It might be surprising that a goalstate of
+ an "unfinished" proof is a theorem. Below
+ we show the internals.
+
+ In general a goalstate is of the form
+
+ A1 \<dots> An \<Longrightarrow> #C
+
+ where the Ai are the open subgoals and C
+ is the theorem to be proved, protected by
+ the constant prop. This constant is usually
+ not visible.
+*}
+
+ML {*
+fun my_print_tac ctxt thm =
+let
+ val _ = tracing (Syntax.string_of_term ctxt (prop_of thm))
+in
+ Seq.single thm
+end
+*}
+
+notation (output) "prop" ("#_" [1000] 1000)
+
+lemma shows "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> A \<and> B"
+apply(tactic {* my_print_tac @{context} *})
+
+apply(rule conjI)
+apply(tactic {* my_print_tac @{context} *})
+
+apply(assumption)
+apply(tactic {* my_print_tac @{context} *})
+
+apply(assumption)
+apply(tactic {* my_print_tac @{context} *})
+done
+
+notation (output) "prop" ("_" [1000] 1000)
+
+text {*
+ There are some simple tactics corresponding
+ to rule, erule and drule (from apply scripts).
+ The examples should be self-explanatory.
+*}
+
+lemma shows "P \<Longrightarrow> P"
+apply(tactic {* atac 1 *})
+oops
+
+lemma shows "P \<and> Q"
+apply(tactic {* resolve_tac [@{thm conjI}] 1 *})
+oops
+
+lemma shows "P \<and> Q \<Longrightarrow> False"
+apply(tactic {* eresolve_tac [@{thm conjE}] 1 *})
+oops
+
+lemma shows "False \<and> True \<Longrightarrow> False"
+apply(tactic {* dresolve_tac [@{thm conjunct2}] 1 *})
+oops
+
+text {*
+ For all sorts of operations on the goalstate there
+ is a corresponding tactic. The following corresponds
+ to "insert".
+*}
+
+lemma shows "True = False"
+apply(tactic {* cut_facts_tac [@{thm True_def}, @{thm False_def}] 1 *})
+oops
+
+text {*
+ Because of schematic variables, theorems need often
+ to be pre-instantiated.
+*}
+
+lemma shows "\<forall>x \<in> A. P x \<Longrightarrow> Q x"
+apply(tactic {* dresolve_tac [@{thm bspec}] 1 *})
+oops
+
+ML {* @{thm disjI1} RS @{thm conjI} *}
+
+text {*
+ Tactic combinators: every tactic (THEN),
+ first applicable tactic (ORELSE).
+*}
+
+ML {*
+val foo_tac' =
+ EVERY' [etac @{thm disjE},
+ rtac @{thm disjI2},
+ atac,
+ rtac @{thm disjI1},
+ atac]
+*}
+
+
+ML {*
+val select_tac =
+ FIRST' [rtac @{thm conjI},
+ rtac @{thm impI},
+ rtac @{thm notI},
+ rtac @{thm allI}, K all_tac]
+*}
+
+text {*
+ The most convenient order to apply tactics to subgoal
+ is in reverse order.
+*}
+
+lemma
+ shows "A \<and> B" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
+apply(tactic {* select_tac 4 *})
+apply(tactic {* select_tac 3 *})
+apply(tactic {* select_tac 2 *})
+apply(tactic {* select_tac 1 *})
+oops
+
+text {*
+ The tactic select_tac can also be written using the
+ combinator TRY.
+*}
+
+ML {*
+val select_tac' =
+ TRY o FIRST' [rtac @{thm conjI},
+ rtac @{thm impI},
+ rtac @{thm notI},
+ rtac @{thm allI}]
+*}
+
+text {*
+ Repeated application of tactics.
+*}
+
+ML {*
+val repeat_sel_tac =
+ REPEAT o CHANGED o select_tac'
+*}
+
+lemma
+ shows "A \<and> (B \<and> C)" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
+apply(tactic {* repeat_sel_tac 4 *})
+apply(tactic {* repeat_sel_tac 3 *})
+apply(tactic {* repeat_sel_tac 2 *})
+apply(tactic {* repeat_sel_tac 1 *})
+oops
+
+text {*
+ Application of a tactics to all new subgoals.
+*}
+
+ML {*
+val repeat_all_new_sel_tac =
+ TRY o REPEAT_ALL_NEW (CHANGED o select_tac')
+*}
+
+lemma
+ shows "A \<and> (B \<and> C)" and "A \<longrightarrow> B \<longrightarrow> C" and " \<forall> x. D x" and "E \<Longrightarrow> F"
+apply(tactic {* repeat_all_new_sel_tac 4 *})
+apply(tactic {* repeat_all_new_sel_tac 3 *})
+apply(tactic {* repeat_all_new_sel_tac 2 *})
+apply(tactic {* repeat_all_new_sel_tac 1 *})
+oops
+
+text {*
+ The tactical DEPTH_SOLVED applies tactics in depth first
+ search including backtracking.
+
+ This can be used to implement a decision procedure for
+ propositional intuitionistic logic inspired by work of
+ Roy Dyckhoff.
+*}
+
+
+lemma impE1:
+ shows "\<lbrakk>A \<longrightarrow> B; A; \<lbrakk>A; B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by iprover
+
+lemma impE2:
+ shows "\<lbrakk>(C \<and> D) \<longrightarrow> B; C \<longrightarrow> (D \<longrightarrow>B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by iprover
+
+lemma impE3:
+ shows "\<lbrakk>(C \<or> D) \<longrightarrow> B; \<lbrakk>C \<longrightarrow> B; D \<longrightarrow> B\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by iprover
+
+lemma impE4:
+ shows "\<lbrakk>(C \<longrightarrow> D) \<longrightarrow> B; D \<longrightarrow> B \<Longrightarrow> C \<longrightarrow> D; B \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by iprover
+
+lemma impE5:
+ shows "\<lbrakk>(C = D) \<longrightarrow> B; (C \<longrightarrow> D) \<longrightarrow> ((D \<longrightarrow> C) \<longrightarrow> B) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
+by iprover
+
+ML {*
+val apply_tac =
+let
+ val intros = [@{thm conjI}, @{thm disjI1}, @{thm disjI2}, @{thm impI}, @{thm iffI}]
+ val elims = [@{thm FalseE}, @{thm conjE}, @{thm disjE}, @{thm iffE},
+ @{thm impE2}, @{thm impE3}, @{thm impE4}, @{thm impE5}, @{thm impE1}]
+in
+ atac
+ ORELSE' resolve_tac intros
+ ORELSE' eresolve_tac elims
+end
+*}
+
+lemma
+ shows "((((P \<longrightarrow> Q) \<longrightarrow> P) \<longrightarrow> P) \<longrightarrow> Q) \<longrightarrow> Q"
+apply(tactic {* (DEPTH_SOLVE o apply_tac) 1 *})
+done
+
+lemma
+ shows "((A \<longrightarrow> B) \<longrightarrow> A) \<longrightarrow> A"
+apply(tactic {* (TRY o DEPTH_SOLVE o apply_tac) 1 *})
+oops
+
+lemma
+ shows "(A \<and> B \<or> ((((A \<longrightarrow> False) \<longrightarrow> False) \<longrightarrow> Q) \<or> (B \<longrightarrow> Q)) \<longrightarrow> Q) \<longrightarrow> Q"
+apply(tactic {* (TRY o DEPTH_SOLVE o apply_tac) 1 *})
+oops
+
+text {*
+ The SUBPROOF tactic gives you full access to
+ all components of a goal state.
+*}
+
+ML {*
+fun subproof_tac {asms, concl, prems, params, context, schematics} =
+let
+ fun out f xs = commas (map (Syntax.string_of_term context o f) xs)
+ val str_of_asms = out term_of asms
+ val str_of_concl = out term_of [concl]
+ val str_of_prems = out prop_of prems
+ val str_of_params = out term_of params
+ val str_of_schms = out term_of (snd schematics)
+ val s = ["params: " ^ str_of_params,
+ "schematics: " ^ str_of_schms,
+ "premises: " ^ str_of_prems,
+ "assumptions: " ^ str_of_asms,
+ "conclusion: " ^ str_of_concl]
+in
+ tracing (cat_lines s); no_tac
+end
+*}
+
+lemma shows "\<And>x y. A x y \<Longrightarrow> B y x \<longrightarrow> C (?z y) x"
+apply(tactic {* SUBPROOF subproof_tac @{context} 1 *})?
+apply(rule impI)
+apply(tactic {* SUBPROOF subproof_tac @{context} 1 *})?
+oops
+
+ML {*
+val my_atac = SUBPROOF (fn {prems, ...} => resolve_tac prems 1)
+*}
+
+lemma shows "\<lbrakk>B x y; A x y; C x y \<rbrakk> \<Longrightarrow> A x y"
+apply(tactic {* my_atac @{context} 1 *})
+done
+
+
+text {*
+ Setting up the goal state using the function
+ Goal.prove. It expects the goal as a term, we
+ have to construct manually.
+*}
+
+ML {*
+open HOLogic
+
+fun P n = @{term "P::nat \<Rightarrow> bool"} $ (mk_number @{typ "nat"} n)
+
+fun rhs 1 = P 1
+ | rhs n = mk_conj (P n, rhs (n - 1))
+
+fun lhs 1 n = mk_imp (mk_eq (P 1, P n), rhs n)
+ | lhs m n = mk_conj (mk_imp
+ (mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
+*}
+
+ML {*
+fun de_bruijn_test ctxt n =
+let
+ val i = 2*n + 1
+ val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
+in
+ Syntax.string_of_term ctxt goal
+ |> writeln
+end;
+
+de_bruijn_test @{context} 1
+*}
+
+ML {*
+fun de_bruijn ctxt n =
+let
+ val i = 2*n+1
+ val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
+in
+ Goal.prove ctxt ["P"] [] goal
+ (fn _ => (DEPTH_SOLVE o apply_tac) 1)
+end
+*}
+
+ML{* de_bruijn @{context} 1 *}
+
+text {*
+ Below is the version which constructs the terms
+ using fixed variables.
+*}
+
+ML {*
+fun de_bruijn ctxt n =
+let
+ val i = 2*n+1
+ val bs = replicate (i+1) "b"
+ val (nbs, ctxt') = Variable.variant_fixes bs ctxt
+ val fbs = map (fn z => Free (z, @{typ "bool"})) nbs
+
+ fun P n = nth fbs n
+
+ fun rhs 0 = @{term "True"}
+ | rhs 1 = P 1
+ | rhs n = mk_conj (P n, rhs (n - 1))
+
+ fun lhs 1 n = mk_imp (mk_eq (P 1, P n), rhs n)
+ | lhs m n = mk_conj (mk_imp
+ (mk_eq (P (m - 1), P m), rhs n), lhs (m - 1) n)
+
+ val goal = mk_Trueprop (mk_imp (lhs i i, rhs i))
+in
+ Goal.prove ctxt' [] [] goal
+ (fn _ => (DEPTH_SOLVE o apply_tac) 1)
+ (*|> singleton (ProofContext.export ctxt' ctxt)*)
+end
+*}
+
+ML {* de_bruijn @{context} 1 *}
+
+
+end
+
+
+
+
+
+
+
+
+
+