diff -r 934d4aed8497 -r 2510e6f7b11c src/CTT/ex/Synthesis.thy --- a/src/CTT/ex/Synthesis.thy Wed Oct 26 17:22:12 2022 +0100 +++ b/src/CTT/ex/Synthesis.thy Wed Oct 26 18:08:44 2022 +0100 @@ -6,40 +6,40 @@ section "Synthesis examples, using a crude form of narrowing" theory Synthesis -imports "../CTT" + imports "../CTT" begin text "discovery of predecessor function" schematic_goal "?a : \pred:?A . Eq(N, pred`0, 0) \ (\n:N. Eq(N, pred ` succ(n), n))" -apply intr -apply eqintr -apply (rule_tac [3] reduction_rls) -apply (rule_tac [5] comp_rls) -apply rew -done + apply intr + apply eqintr + apply (rule_tac [3] reduction_rls) + apply (rule_tac [5] comp_rls) + apply rew + done text "the function fst as an element of a function type" schematic_goal [folded basic_defs]: "A type \ ?a: \f:?B . \i:A. \j:A. Eq(A, f ` , i)" -apply intr -apply eqintr -apply (rule_tac [2] reduction_rls) -apply (rule_tac [4] comp_rls) -apply typechk -txt "now put in A everywhere" -apply assumption+ -done + apply intr + apply eqintr + apply (rule_tac [2] reduction_rls) + apply (rule_tac [4] comp_rls) + apply typechk + txt "now put in A everywhere" + apply assumption+ + done text "An interesting use of the eliminator, when" -(*The early implementation of unification caused non-rigid path in occur check + (*The early implementation of unification caused non-rigid path in occur check See following example.*) schematic_goal "?a : \i:N. Eq(?A, ?b(inl(i)), <0 , i>) \ Eq(?A, ?b(inr(i)), )" -apply intr -apply eqintr -apply (rule comp_rls) -apply rew -done + apply intr + apply eqintr + apply (rule comp_rls) + apply rew + done (*Here we allow the type to depend on i. This prevents the cycle in the first unification (no longer needed). @@ -47,58 +47,58 @@ Simpler still: make ?A into a constant type N \ N.*) schematic_goal "?a : \i:N. Eq(?A(i), ?b(inl(i)), <0 , i>) \ Eq(?A(i), ?b(inr(i)), )" -oops + oops -text "A tricky combination of when and split" -(*Now handled easily, but caused great problems once*) + text "A tricky combination of when and split" + (*Now handled easily, but caused great problems once*) schematic_goal [folded basic_defs]: "?a : \i:N. \j:N. Eq(?A, ?b(inl()), i) \ Eq(?A, ?b(inr()), j)" -apply intr -apply eqintr -apply (rule PlusC_inl [THEN trans_elem]) -apply (rule_tac [4] comp_rls) -apply (rule_tac [7] reduction_rls) -apply (rule_tac [10] comp_rls) -apply typechk -done + apply intr + apply eqintr + apply (rule PlusC_inl [THEN trans_elem]) + apply (rule_tac [4] comp_rls) + apply (rule_tac [7] reduction_rls) + apply (rule_tac [10] comp_rls) + apply typechk + done (*similar but allows the type to depend on i and j*) schematic_goal "?a : \i:N. \j:N. Eq(?A(i,j), ?b(inl()), i) \ Eq(?A(i,j), ?b(inr()), j)" -oops + oops (*similar but specifying the type N simplifies the unification problems*) schematic_goal "?a : \i:N. \j:N. Eq(N, ?b(inl()), i) \ Eq(N, ?b(inr()), j)" -oops + oops -text "Deriving the addition operator" + text "Deriving the addition operator" schematic_goal [folded arith_defs]: "?c : \n:N. Eq(N, ?f(0,n), n) \ (\m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))" -apply intr -apply eqintr -apply (rule comp_rls) -apply rew -done + apply intr + apply eqintr + apply (rule comp_rls) + apply rew + done text "The addition function -- using explicit lambdas" schematic_goal [folded arith_defs]: "?c : \plus : ?A . \x:N. Eq(N, plus`0`x, x) \ (\y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))" -apply intr -apply eqintr -apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3") -apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4") -apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3") -apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4") -apply (rule_tac [3] p = "y" in NC_succ) - (** by (resolve_tac @{context} comp_rls 3); caused excessive branching **) -apply rew -done + apply intr + apply eqintr + apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3") + apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4") + apply (tactic "resolve_tac \<^context> [TSimp.split_eqn] 3") + apply (tactic "SELECT_GOAL (rew_tac \<^context> []) 4") + apply (rule_tac [3] p = "y" in NC_succ) + (** by (resolve_tac @{context} comp_rls 3); caused excessive branching **) + apply rew + done end