From 01dedbf48ed5dac009134b04aaa82aa4eff448e7 Mon Sep 17 00:00:00 2001 From: Ra Date: Thu, 12 Dec 2024 20:09:28 +1100 Subject: [PATCH 1/2] Pedersen commitment scheme --- .../Crypto/pedersenCommitment/Holmakefile | 4 + .../pedersenCommitmentScript.sml | 220 ++++++++++++++++++ 2 files changed, 224 insertions(+) create mode 100755 examples/Crypto/pedersenCommitment/Holmakefile create mode 100755 examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml diff --git a/examples/Crypto/pedersenCommitment/Holmakefile b/examples/Crypto/pedersenCommitment/Holmakefile new file mode 100755 index 0000000000..ff60aa1f91 --- /dev/null +++ b/examples/Crypto/pedersenCommitment/Holmakefile @@ -0,0 +1,4 @@ +INCLUDES = $(HOLDIR)/examples/algebra/field \ + $(HOLDIR)/src/algebra/construction \ + $(HOLDIR)/src/algebra/construction \ + $(HOLDIR)/src/num/theories \ No newline at end of file diff --git a/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml b/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml new file mode 100755 index 0000000000..6d45189eee --- /dev/null +++ b/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml @@ -0,0 +1,220 @@ + +(* ------------------------------------------------------------------------- *) +(* Theory: Pedersen Commitment Scheme *) +(* ------------------------------------------------------------------------- *) + +open HolKernel boolLib bossLib Parse; +open numTheory jcLib monoidTheory fieldTheory fieldInstancesTheory; +open groupTheory arithmeticTheory dividesTheory numeralTheory ringTheory; + +(* Declare new theory at start *) +val _ = new_theory "pedersenCommitment"; + +(* Overload operations to avoid ambiguity *) +Overload "_0_" = “(GF q).sum.id”; +Overload "_1_" = “(GF q).prod.id”; +Overload "_+_" = “(GF q).sum.op”; val _ = set_fixity "_+_" (Infixl 500); +Overload "_-_" = “ring_sub (GF q)”; val _ = set_fixity "_-_" (Infixl 500); +Overload "_*_" = “(GF q).prod.op”; val _ = set_fixity "_*_" (Infixl 600); +Overload "_/_" = “field_div (GF q)”; val _ = set_fixity "_/_" (Infixl 600); + +(* Theorems *) + +val field_mult_solve_eqn = store_thm( + "field_mult_solve_eqn", + “!r:'a field. Field r ==> + !x y z. x IN R /\ y IN R /\ z IN R /\ y <> #0 ==> + (x = z * y <=> x / y = z)”, + rw[EQ_IMP_THM] >- + (simp[field_nonzero_eq, field_mult_assoc, field_mult_rinv]) >> + simp[field_nonzero_eq, field_mult_assoc, field_mult_linv]); + +val GF_mult_solve_eqn = store_thm( + "GF_mult_solve_eqn", + “!q. prime q ==> + !d m i. m IN (GF q).carrier /\ + d IN (GF q).carrier /\ + i IN (GF q).carrier /\ m <> _0_ ==> + (d = i _*_ m <=> d _/_ m = i)”, + simp[GF_field, field_mult_solve_eqn]); + +val GF_sub_not_eq_zero = store_thm( + "GF_sub_not_eq_zero", + “∀q. prime q ⇒ + ∀m1 m2. + m1 IN (GF q).carrier ∧ + m2 IN (GF q).carrier ∧ + m2 ≠ m1 ⇒ + m2 _-_ m1 ≠ _0_”, + rpt strip_tac >> + ‘Field (GF q)’ by rw[GF_field] >> + metis_tac[field_sub_eq_zero]); + +val GF_mult_rsub = store_thm( + "GF_mult_rsub", + “∀q. prime q ⇒ + ∀m1 m2 i. + m1 IN (GF q).carrier ∧ + m2 IN (GF q).carrier ∧ + i IN (GF q).carrier ⇒ + (i _*_ m2) _-_ (i _*_ m1) = i _*_ (m2 _-_ m1)”, + rpt strip_tac >> + ‘Field (GF q)’ by rw[GF_field] >> + metis_tac[field_mult_rsub]); + +val field_mult_not_zero = store_thm( + "field_mult_not_zero", + “!r:'a field. Field r ==> + !d m i. d IN R ∧ m IN R ∧ i IN R ∧ d = i * m ∧ i ≠ #0 ∧ m ≠ #0 ==> (d ≠ #0)”, + rpt strip_tac >> + rw[] >> + ‘i = #0 ∨ m = #0’ by metis_tac[field_mult_eq_zero]); + +val GF_mult_not_zero = store_thm( + "GF_mult_not_zero", + “∀q. prime q ⇒ + ∀d m i. + m IN (GF q).carrier ∧ + d IN (GF q).carrier ∧ + i IN (GF q).carrier ∧ + d = i _*_ m ∧ + i ≠ _0_ ∧ + m ≠ _0_ ⇒ d ≠ _0_”, + rpt strip_tac >> + ‘Field (GF q)’ by rw[GF_field] >> + metis_tac[field_mult_not_zero]); + +val GF_add_sub_identity = store_thm( + "GF_add_sub_identity", + “∀q. prime q ⇒ + ∀x y z t. + x IN (GF q).carrier ∧ + y IN (GF q).carrier ∧ + z IN (GF q).carrier ∧ + t IN (GF q).carrier ∧ + x _+_ y = z _+_ t ⇒ + x _-_ z = t _-_ y”, + rpt strip_tac >> + ‘Field (GF q)’ by metis_tac[GF_field] >> + metis_tac[field_add_sub_identity]); + +val GF_diff_mult_solve = store_thm( + "GF_diff_mult_solve", + “∀q. prime q ⇒ + ∀d1 d2 m1 m2 i. + d1 IN (GF q).carrier ∧ + d2 IN (GF q).carrier ∧ + m1 IN (GF q).carrier ∧ + m2 IN (GF q).carrier ∧ + i IN (GF q).carrier ∧ + m2 ≠ m1 ∧ + d1 _-_ d2 = i _*_ (m2 _-_ m1) ⇒ + (d1 _-_ d2) _/_ (m2 _-_ m1) = i”, + rpt strip_tac >> + ‘Field (GF q)’ by rw[GF_field] >> + ‘m2 _-_ m1 ≠ _0_’ by metis_tac[GF_sub_not_eq_zero] >> + ‘(d1 _-_ d2) IN (GF q).carrier ∧ (m2 _-_ m1) IN (GF q).carrier’ by rw[] >> + metis_tac[GF_mult_solve_eqn]); + +(* Definitions *) +Definition commit_def: + commit (g: 'a group) (h: 'a) d k m = g.op (monoid_exp g h d) (monoid_exp g k m) +End + +Definition verify_def: + verify (g: 'a group) (h: 'a) c d k m ⇔ commit g h d k m = c +End + +(* Theorems involving commit and verify *) +val GF_pedersen_binding = store_thm( + "GF_pedersen_binding", + “∀q. prime q ⇒ + ∀m1 m2 d1 d2. + m1 IN (GF q).carrier ∧ + m2 IN (GF q).carrier ∧ + d1 IN (GF q).carrier ∧ + d2 IN (GF q).carrier ∧ + (m1 ≠ m2) ⇒ + ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ + ∀k. k IN G ∧ + verify g (cyclic_gen g) (commit g (cyclic_gen g) d1 k m1) d2 k m2 + ⇒ + (d1 _-_ d2) _/_ (m2 _-_ m1) = cyclic_index g k”, + simp[commit_def, verify_def, Excl "ring_sub_def", Excl "field_div_def"] >> + rpt strip_tac >> + qabbrev_tac ‘h = cyclic_gen g’ >> + ‘Group g’ by metis_tac[cyclic_def] >> + ‘cyclic_gen g IN G’ by metis_tac[cyclic_gen_element] >> + ‘(ord (cyclic_gen g)) = CARD G’ by metis_tac[cyclic_gen_order] >> + ‘∃i. i < q ∧ k = monoid_exp g h i’ by metis_tac[cyclic_element_by_gen] >> + ‘cyclic_index g k = i’ by metis_tac[finite_cyclic_index_unique] >> + ‘h IN G’ by fs[cyclic_gen_element, Abbr`h`] >> + ‘i IN (GF q).carrier’ by rw[GF_eval] >> + fs[GSYM group_exp_mult, Excl "monoid_exp_mult", GSYM group_exp_add, Excl "monoid_exp_add", + Excl "ring_sub_def", Excl "field_div_def"] >> + ‘0 < q’ by metis_tac[NOT_PRIME_0, NOT_ZERO] >> + ‘monoid_exp g h ((d2 + i * m2) MOD q) = monoid_exp g h ((d1 + i * m1) MOD q)’ by metis_tac[group_exp_mod] >> + qabbrev_tac ‘x2 = monoid_exp g h ((d2 + i * m2) MOD q)’ >> + qabbrev_tac ‘x1 = monoid_exp g h ((d1 + i * m1) MOD q)’ >> + ‘x2 IN G ∧ x1 IN G’ by fs[Abbr`x2`, Abbr`x1`] >> + ‘((d2 + i * m2) MOD q) = cyclic_index g x2 ∧ ((d1 + i * m1) MOD q) = cyclic_index g x1’ + by metis_tac[finite_cyclic_index_unique, MOD_LESS] >> + ‘((d2 + i * m2) MOD q) = ((d1 + i * m1) MOD q)’ by rw[] >> + ‘d2 _+_ i _*_ m2 = d1 _+_ i _*_ m1’ by metis_tac[MOD_PLUS, MOD_MOD, GF_eval] >> + ‘Field (GF q)’ by rw[GF_field] >> + ‘d1 _-_ d2 = (i _*_ m2) _-_ (i _*_ m1)’ by rw[GF_add_sub_identity, GF_eval] >> + ‘d1 _-_ d2 = i _*_ (m2 _-_ m1)’ by rw[GF_mult_rsub, GF_eval] >> + prove_tac[GF_diff_mult_solve]); + +Theorem GF_pedersen_binding_prime = GF_pedersen_binding |> SIMP_RULE (srw_ss()) [verify_def]; + +val GF_pedersen_hiding = store_thm( + "GF_pedersen_hiding", + “∀q. prime q ⇒ + ∀m. m IN (GF q).carrier ⇒ + ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ + ∀c k. c IN G ∧ k IN G ⇒ + ∃d. d IN (GF q).carrier ∧ + c = commit g (cyclic_gen g) d k m”, + simp[commit_def] >> + rw[] >> + ‘∃i. i < CARD G ∧ c = monoid_exp g (cyclic_gen g) i’ by rw[cyclic_element_by_gen] >> + ‘∃j. j < CARD G ∧ k = monoid_exp g (cyclic_gen g) j’ by rw[cyclic_element_by_gen] >> + ‘Group g’ by metis_tac[cyclic_def] >> + rw[Once EQ_SYM_EQ] >> + rw[group_lsolve, cyclic_gen_element, cyclic_gen_order] >> + ‘monoid_exp g (monoid_exp g (cyclic_gen g) j) m IN G’ by rw[] >> + ‘monoid_inv g (monoid_exp g (monoid_exp g (cyclic_gen g) j) m) IN G’ by rw[group_inv_element] >> + ‘(g.op (monoid_exp g (cyclic_gen g) i) (monoid_inv g (monoid_exp g (monoid_exp g (cyclic_gen g) j) m))) IN G’ + by rw[group_op_element] >> + qabbrev_tac ‘x = (g.op (monoid_exp g (cyclic_gen g) i) (monoid_inv g (monoid_exp g (monoid_exp g (cyclic_gen g) j) m)))’ >> + ‘∃d. d < CARD G ∧ x = monoid_exp g (cyclic_gen g) d’ by rw[cyclic_element_by_gen] >> + ‘ord (cyclic_gen g) = CARD G’ by rw[cyclic_gen_order] >> + ‘d ∈ (GF (CARD G)).carrier’ by rw[GF_eval] >> + qexists_tac ‘d’ >> + rw[]); + +val ZN_pedersen_hiding = store_thm( + "ZN_pedersen_hiding", + “∀q. prime q ⇒ + ∀m. m IN (ZN q).carrier ⇒ + ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ + ∀c k. c IN G ∧ k IN G ⇒ + ∃d. d IN (ZN q).carrier ∧ + c = commit g (cyclic_gen g) d k m”, + simp[commit_def] >> + rw[] >> + ‘∃i. i < CARD G ∧ c = monoid_exp g (cyclic_gen g) i’ by rw[cyclic_element_by_gen] >> + ‘∃j. j < CARD G ∧ k = monoid_exp g (cyclic_gen g) j’ by rw[cyclic_element_by_gen] >> + ‘Group g’ by metis_tac[cyclic_def] >> + rw[Once EQ_SYM_EQ] >> + rw[ZN_def] >> + rw[cyclic_gen_order] >> + rw[group_lsolve, cyclic_gen_element, cyclic_gen_order] >> + rw[Once EQ_SYM_EQ] >> + irule cyclic_element_by_gen >> + simp[]); + +Theorem GF_pedersen_hiding_prime = GF_pedersen_hiding |> SIMP_RULE (srw_ss()) [verify_def]; + +val _ = export_theory(); From 6f5ba93828d16a464693f2c8c7693122950aa750 Mon Sep 17 00:00:00 2001 From: Ra Date: Sun, 12 Jan 2025 11:31:12 +1100 Subject: [PATCH 2/2] sigma protocol --- examples/AKS/compute/files.txt | 35 - examples/AKS/machine/files.txt | 108 - examples/AKS/theories/files.txt | 66 - examples/Crypto/Keccak/keccakScript.sml | 2127 +++++- examples/Crypto/RSA/binomialScript.sml | 3 - examples/Crypto/RSA/congruentScript.sml | 3 - examples/Crypto/RSA/fermatScript.sml | 2 - examples/Crypto/RSA/powerScript.sml | 3 - examples/Crypto/RSA/rsaScript.sml | 4 - examples/Crypto/RSA/summationScript.sml | 2 - .../Crypto/pedersenCommitment/Holmakefile | 3 +- .../pedersenCommitmentScript.sml | 40 +- examples/Crypto/sigmaProtocol/Holmakefile | 6 + .../sigmaProtocol/sigmaProtocolScript.sml | 1339 ++++ examples/algebra/README.md | 47 +- examples/algebra/field/files.txt | 62 - examples/algebra/finitefield/files.txt | 84 - examples/algebra/linear/files.txt | 21 - examples/algebra/polynomial/files.txt | 209 - examples/arm/arm6-verification/armScript.sml | 3 - .../arm6-verification/io_onestepScript.sml | 46 +- .../arm/arm6-verification/lemmasScript.sml | 13 +- .../arm/arm6-verification/my_listScript.sml | 26 +- .../arm/arm6-verification/onestepScript.sml | 28 +- examples/arm/v4/arm_evalScript.sml | 13 +- examples/diningcryptos/extra_stringScript.sml | 299 +- examples/fermat/count/files.txt | 36 - examples/fermat/little/files.txt | 68 - examples/fermat/twosq/README.md | 6 +- examples/fermat/twosq/files.txt | 72 - .../formal-languages/FormalLangScript.sml | 6 + examples/formal-languages/regular/Holmakefile | 3 +- .../regular/regularScript.sml | 620 +- .../common/temporal_stateScript.sml | 3 - examples/lambda/barendregt/boehmScript.sml | 5897 +++++++++-------- .../barendregt/head_reductionScript.sml | 15 + .../barendregt/lameta_completeScript.sml | 381 +- examples/lambda/basics/basic_swapScript.sml | 10 + examples/lambda/basics/termScript.sml | 28 + .../logic/propositional_logic/Holmakefile | 2 +- .../arm_cheney_allocScript.sml | 2 - .../arm_cheney_gcScript.sml | 3 - .../arm_improved_gcScript.sml | 2 - .../garbage-collectors/cheney_allocScript.sml | 2 - .../garbage-collectors/cheney_gcScript.sml | 3 +- .../garbage-collectors/improved_gcScript.sml | 2 - .../garbage-collectors/lisp_gcScript.sml | 2 - .../hoare-triple/tailrecScript.sml | 4 - .../arm/prog_armLib.sml | 4 - .../arm/prog_armScript.sml | 3 - .../ppc/prog_ppcScript.sml | 4 - .../x86/prog_x86Lib.sml | 5 - .../x86/prog_x86Script.sml | 4 - .../x86_64/prog_x64Lib.sml | 3 - .../x86_64/prog_x64Script.sml | 3 - .../x86_64/prog_x64_extraScript.sml | 2 - .../just-in-time/jit_codegenScript.sml | 4 - .../just-in-time/jit_incrementalScript.sml | 4 - .../just-in-time/jit_inputScript.sml | 2 - .../just-in-time/jit_opsScript.sml | 4 - .../machine-code/lisp/lisp_equalScript.sml | 2 - examples/machine-code/lisp/lisp_invScript.sml | 2 - examples/machine-code/lisp/lisp_opsScript.sml | 2 - .../machine-code/lisp/lisp_parseScript.sml | 2 - .../machine-code/lisp/lisp_printScript.sml | 2 - .../machine-code/lisp/lisp_typeScript.sml | 2 - examples/miller/formalize/orderScript.sml | 73 +- examples/miller/formalize/sequenceScript.sml | 102 +- examples/pgcl/src/relScript.sml | 7 - examples/simple_complexity/lib/files.txt | 15 - examples/simple_complexity/loop/files.txt | 44 - .../bytecode/lisp_alt_semanticsScript.sml | 2 - .../bytecode/lisp_bytecodeScript.sml | 2 - .../bytecode/lisp_compilerScript.sml | 2 - .../extract/lisp_extractScript.sml | 2 - .../extract/lisp_synthesis_demoScript.sml | 3 - .../garbage-collector/lisp_consScript.sml | 2 - .../garbage-collector/stop_and_copyScript.sml | 2 - .../implementation/lisp_bigopsScript.sml | 2 - .../lisp_bytecode_stepScript.sml | 2 - .../implementation/lisp_codegenScript.sml | 2 - .../implementation/lisp_compiler_opScript.sml | 2 - .../implementation/lisp_correctnessScript.sml | 2 - .../implementation/lisp_equalScript.sml | 2 - .../implementation/lisp_invScript.sml | 2 - .../implementation/lisp_opsScript.sml | 2 - .../implementation/lisp_symbolsScript.sml | 2 - .../lisp-runtime/parse/lisp_parseScript.sml | 2 - .../milawa-prover/milawa_coreScript.sml | 2 - .../milawa-prover/milawa_defsScript.sml | 2 - .../milawa-prover/milawa_execScript.sml | 4 - .../milawa-prover/milawa_initScript.sml | 2 - .../milawa-prover/milawa_logicScript.sml | 3 - .../milawa-prover/milawa_ordinalScript.sml | 2 - .../milawa-prover/milawa_proofpScript.sml | 2 - .../soundness-thm/milawa_soundnessScript.sml | 2 - .../Portable.remove_external_wspace.doc | 25 + help/Docfiles/Portable.remove_wspace.doc | 3 + src/1/ThmAttribute.sig | 21 +- src/1/ThmAttribute.sml | 89 +- src/1/ThmSetData.sml | 37 +- src/1/boolLib.sml | 37 +- src/1/boolSyntax.sml | 36 +- src/1/selftest.sml | 26 + src/1/theory_tests/Holmakefile | 3 + src/1/theory_tests/gh1370Script.sml | 15 + src/1/theory_tests/gh1370childScript.sml | 11 + src/bag/primeFactorScript.sml | 1 - src/bool/boolScript.sml | 1 - src/boss/prove_base_assumsScript.sml | 227 +- src/coalgebras/ltreeScript.sml | 2 +- src/coalgebras/pathScript.sml | 3 - src/combin/combinScript.sml | 89 +- src/compute/src/computeLib.sig | 2 + src/compute/src/computeLib.sml | 77 +- src/float/ieeeScript.sml | 2 - src/integer/DeepSyntaxScript.sml | 3 - src/integer/int_arithScript.sml | 7 +- src/integer/int_bitwiseScript.sml | 3 - src/list/src/listScript.sml | 10 + src/list/src/numposrepScript.sml | 84 + src/list/src/rich_listScript.sml | 217 +- .../more_monads/state_transformerScript.sml | 3 - src/n-bit/wordsScript.sml | 218 + src/num/extra_theories/bitScript.sml | 57 + src/num/termination/TotalDefn.sml | 49 +- src/num/theories/arithmeticScript.sml | 48 + .../cv_compute/automation/cv_primScript.sml | 13 +- .../cv_compute/automation/cv_typeScript.sml | 19 +- src/parallel_builds/core/Holmakefile | 3 +- src/parse/Preterm.sml | 2 + src/parse/term_pp.sml | 2 + src/portableML/HOLsexp.sig | 2 + src/portableML/HOLsexp.sml | 16 +- src/portableML/Portable.sig | 1 + src/portableML/Portable.sml | 6 + src/postkernel/TheoryReader.sml | 15 +- src/postkernel/ThyDataSexp.sig | 1 + src/quotient/examples/finite_setScript.sml | 3 - src/quotient/examples/lambda/betaScript.sml | 4 - src/quotient/examples/lambda/etaScript.sml | 4 - .../examples/lambda/reductionScript.sml | 4 - src/quotient/examples/lambda/termScript.sml | 2 - .../examples/lambda/variableScript.sml | 3 - src/quotient/examples/more_listScript.sml | 3 - src/quotient/examples/sigma/alphaScript.sml | 3 - src/quotient/examples/sigma/objectScript.sml | 3 - .../examples/sigma/reductionScript.sml | 3 - .../examples/sigma/variableScript.sml | 3 - src/tfl/examples/sorting/partitionScript.sml | 4 - src/tfl/examples/sorting/permScript.sml | 3 - src/tfl/examples/sorting/sortingScript.sml | 3 - src/transfer/transferLib.sml | 3 +- tools/Holmake/HM_Core_Cline.sig | 1 + tools/Holmake/HM_Core_Cline.sml | 27 +- tools/Holmake/HM_DepGraph.sig | 1 + tools/Holmake/HM_DepGraph.sml | 52 + tools/Holmake/HolLex | 2 +- tools/Holmake/Holmake.sml | 29 +- tools/Holmake/Holmake_tools.sig | 5 + tools/Holmake/Holmake_tools.sml | 13 + .../Holmake/tests/quote-filter/expected-mosml | 9 + .../Holmake/tests/quote-filter/expected-poly | 9 + tools/Holmake/tests/quote-filter/input | 8 + 164 files changed, 9088 insertions(+), 4710 deletions(-) delete mode 100644 examples/AKS/compute/files.txt delete mode 100644 examples/AKS/machine/files.txt delete mode 100644 examples/AKS/theories/files.txt create mode 100755 examples/Crypto/sigmaProtocol/Holmakefile create mode 100644 examples/Crypto/sigmaProtocol/sigmaProtocolScript.sml delete mode 100644 examples/algebra/field/files.txt delete mode 100644 examples/algebra/finitefield/files.txt delete mode 100644 examples/algebra/linear/files.txt delete mode 100644 examples/algebra/polynomial/files.txt delete mode 100644 examples/fermat/count/files.txt delete mode 100644 examples/fermat/little/files.txt delete mode 100644 examples/fermat/twosq/files.txt delete mode 100644 examples/simple_complexity/lib/files.txt delete mode 100644 examples/simple_complexity/loop/files.txt create mode 100644 help/Docfiles/Portable.remove_external_wspace.doc create mode 100644 src/1/theory_tests/Holmakefile create mode 100644 src/1/theory_tests/gh1370Script.sml create mode 100644 src/1/theory_tests/gh1370childScript.sml diff --git a/examples/AKS/compute/files.txt b/examples/AKS/compute/files.txt deleted file mode 100644 index e765ce5818..0000000000 --- a/examples/AKS/compute/files.txt +++ /dev/null @@ -1,35 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of AKS Computation Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2016 *) -(* ------------------------------------------------------------------------- *) - -0 computeBasic -- computations of exp, root. -* Euler -* Gauss -* logPower -* while - -1 computeOrder -- computation of modular multiplicative order. -* groupOrder -* 0 computeBasic - -2 computeParam -- computation of AKS parameter. -* logPower -* 1 computeOrder - -2 computePoly -- polynomial computations with modulus unity. -* 0 computeBasic -* 1 computeOrder - -3 computeRing -- modulo polynomial computations in ring (ZN n). -* 0 computeBasic -* 1 computeOrder -* 2 computePoly - -4 computeAKS -- polynomial checks and all part of the AKS algorithm. -* 0 computeBasic -* 1 computeOrder -* 2 computePoly -* 3 computeRing diff --git a/examples/AKS/machine/files.txt b/examples/AKS/machine/files.txt deleted file mode 100644 index dae9d4e6de..0000000000 --- a/examples/AKS/machine/files.txt +++ /dev/null @@ -1,108 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of AKS Machine Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2016 *) -(* ------------------------------------------------------------------------- *) - -0 countMonad -- a monad keeping track of computational values and counts. -* pair -* option -* errorStateMonad - -1 countMacro -- macro operations as subroutines in monadic style. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* loopDivide -* loopMultiply -* loopList -* 0 countMonad - -2 countBasic -- basic arithmetic functions in monadic style. -* logroot -* logPower -* bitsize -* complexity -* loopIncrease -* loopDecrease -* loopDivide -* loopMultiply -* 0 countMonad -* 1 countMacro - -2 countModulo -- modulo arithmetic in monadic style. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* loopDivide -* loopMultiply -* 0 countMonad -* 1 countMacro - -3 countPower -- exponentiation, root extraction and power free test. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* loopDivide -* 0 countMonad -* 1 countMacro -* 2 countBasic - -3 countOrder -- modular exponentiation and multiplicative order. -* ringInstances -* bitsize -* complexity -* loopIncrease -* loopDecrease -* 0 countMonad -* 1 countMacro -* 2 countModulo - -3 countPoly -- polynomial algorithms for AKS algorithm. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* loopDivide -* loopList -* 0 countMonad -* 1 countMacro -* 2 countModulo - -4 countPrime -- traditional primality testing algorithm. -* bitsize -* complexity -* loopIncrease -* primes -* 0 countMonad -* 1 countMacro -* 2 countBasic -* 3 countPower - -4 countParam -- parameter search for AKS algorithm. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* 0 countMonad -* 1 countMacro -* 2 countBasic -* 3 countPower -* 3 countOrder - -5 countAKS -- combine with polynomial checks for AKS algorithm. -* bitsize -* complexity -* loopIncrease -* loopDecrease -* 0 countMonad -* 1 countMacro -* 2 countBasic -* 3 countPower -* 3 countOrder -* 3 countPoly -* 4 countParam diff --git a/examples/AKS/theories/files.txt b/examples/AKS/theories/files.txt deleted file mode 100644 index 833e89c450..0000000000 --- a/examples/AKS/theories/files.txt +++ /dev/null @@ -1,66 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of AKS Theories Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2014 *) -(* ------------------------------------------------------------------------- *) - -0 AKSintro -- introspective relation for special polynomials and special exponents. -* ffBasic -* ffAdvanced -* ffPoly -* ffUnity -* ffConjugate -* ffExist -* computeRing (* for overloads *) - -1 AKSshift -- introspective shifting: from Ring (ZN n) to Ring (ZN p). -* polyFieldModulo -* polyMap -* 0 AKSintro - -1 AKSsets -- introspective sets, for exponents and polynomials. -* Gauss -* 0 AKSintro - -2 AKSmaps -- mappings between introspective sets and their shadows. -* logPower -* 0 AKSintro -* 1 AKSsets - -3 AKStheorem -- the AKS Main Theorem, with parameter k be prime. -* Euler -* logPower -* computeAKS -* computeRing -* 0 AKSintro -* 1 AKSsets -* 1 AKSshift -* 2 AKSmaps - -4 AKSrevised -- the AKS Main Theorem, with parameter k not required to be prime. -* Gauss -* logPower -* computeAKS -* 0 AKSintro -* 1 AKSsets -* 1 AKSshift -* 2 AKSmaps -* 3 AKStheorem - -5 AKSimproved -- the AKS Main Theorem, with bounds improved by cofactor. -* 0 AKSintro -* 1 AKSsets -* 1 AKSshift -* 2 AKSmaps -* 3 AKStheorem -* 4 AKSrevised - -6 AKSclean -- a clean version of the AKS Main Theorem -* 0 AKSintro -* 1 AKSsets -* 1 AKSshift -* 2 AKSmaps -* 3 AKStheorem -* 4 AKSrevised -* 5 AKSimproved diff --git a/examples/Crypto/Keccak/keccakScript.sml b/examples/Crypto/Keccak/keccakScript.sml index 45f61c5f67..598322f312 100644 --- a/examples/Crypto/Keccak/keccakScript.sml +++ b/examples/Crypto/Keccak/keccakScript.sml @@ -1,8 +1,10 @@ -open HolKernel Parse boolLib bossLib dep_rewrite +open HolKernel Parse boolLib bossLib dep_rewrite blastLib bitLib reduceLib combinLib optionLib sptreeLib wordsLib computeLib; open optionTheory pairTheory arithmeticTheory combinTheory listTheory rich_listTheory whileTheory bitTheory dividesTheory wordsTheory - logrootTheory sptreeTheory; + indexedListsTheory numposrepTheory numeral_bitTheory + bitstringTheory logrootTheory byteTheory sptreeTheory; +open cv_transLib cvTheory cv_stdTheory; (* The SHA-3 Standard: https://doi.org/10.6028/NIST.FIPS.202 *) @@ -10,6 +12,110 @@ val _ = new_theory "keccak"; val _ = numLib.temp_prefer_num(); +Theorem LENGTH_PAD_RIGHT_0_8_word_to_bin_list[simp]: + LENGTH (PAD_RIGHT 0 8 (word_to_bin_list (w: word8))) = 8 +Proof + rw[word_to_bin_list_def, wordsTheory.w2l_def, PAD_RIGHT, LENGTH_n2l, LOG2_def] + \\ Cases_on`w` \\ gs[] + \\ rewrite_tac[Once ADD_SYM] + \\ irule SUB_ADD + \\ `n < 2 ** 8` by simp[] + \\ gs[LT_EXP_LOG] +QED + +Theorem LENGTH_PAD_RIGHT_0_64_word_to_bin_list[simp]: + LENGTH (PAD_RIGHT 0 64 (word_to_bin_list (w: word64))) = 64 +Proof + rw[word_to_bin_list_def, wordsTheory.w2l_def, PAD_RIGHT, LENGTH_n2l, LOG2_def] + \\ Cases_on`w` \\ gs[] + \\ rewrite_tac[Once ADD_SYM] + \\ irule SUB_ADD + \\ `n < 2 ** 64` by simp[] + \\ gs[LT_EXP_LOG] +QED + +Theorem concat_word_list_bytes_to_64: + LENGTH (ls: word8 list) = 8 ⇒ + concat_word_list ls : word64 = + word_from_bin_list (FLAT (MAP (PAD_RIGHT 0 8 o word_to_bin_list) ls)) +Proof + simp[LENGTH_EQ_NUM_compute, PULL_EXISTS] + \\ qx_genl_tac[`b1`,`b2`,`b3`,`b4`,`b5`,`b6`,`b7`,`b8`] + \\ rw[concat_word_list_def] + \\ rw[word_from_bin_list_def, l2w_def] + \\ rw[l2n_APPEND] + \\ rw[word_to_bin_list_def, wordsTheory.w2l_def] + \\ simp[l2n_n2l] + \\ simp[GSYM word_add_n2w] + \\ map_every Cases_on [`b1`,`b2`,`b3`,`b4`,`b5`,`b6`,`b7`,`b8`] + \\ gs[] + \\ simp[w2w_def] + \\ simp[GSYM word_mul_n2w, word_or_n2w] + \\ rw(List.tabulate(8, fn i => + WORD_MUL_LSL |> CONV_RULE SWAP_FORALL_CONV + |> Q.SPEC`8 * ^(numSyntax.term_of_int i)` + |> GSYM |> SIMP_RULE std_ss [])) + \\ qmatch_goalsub_abbrev_tac `a1 + a2 + a3 + a4 + a5 + a6 + a7 + a8` + \\ `∀i m. m < 256 ∧ 8 ≤ i ⇒ ¬BIT i m` + by ( + ntac 3 strip_tac + \\ Cases_on`m=0` >- simp[] + \\ irule NOT_BIT_GT_LOG2 + \\ qspecl_then[`m`,`8`]mp_tac LT_TWOEXP + \\ simp[] ) + \\ DEP_REWRITE_TAC[WORD_ADD_OR] + \\ rewrite_tac[GSYM WORD_OR_ASSOC] + \\ unabbrev_all_tac + \\ blastLib.BBLAST_TAC + \\ simp[] +QED + +Theorem cv_LEN_add: + cv_LEN cv (Num n) = + cv_add (Num n) (cv_LEN cv (Num 0)) +Proof + qid_spec_tac`n` + \\ Induct_on`cv` + >- ( rw[Once cv_LEN_def] \\ rw[Once cv_LEN_def] ) + \\ rewrite_tac[Once cv_LEN_def] + \\ simp_tac (srw_ss()) [] + \\ gen_tac + \\ simp_tac std_ss [Once cv_LEN_def, SimpRHS] + \\ simp_tac (srw_ss()) [] + \\ first_assum(qspec_then`n + 1`SUBST1_TAC) + \\ first_assum(qspec_then`1`SUBST1_TAC) + \\ qmatch_goalsub_abbrev_tac`cv_add _ p` + \\ Cases_on`p` + \\ simp[] +QED + +Theorem DIV2_SBIT: + DIV2 (SBIT b n) = if n = 0 then 0 else SBIT b (n - 1) +Proof + qspecl_then[`b`,`n`,`1`]mp_tac SBIT_DIV + \\ simp[DIV2_def] + \\ Cases_on`1 < n` \\ gs[] + \\ Cases_on`n=0` \\ gs[SBIT_def] + \\ rw[] + \\ `n=1` by gs[] + \\ gvs[] +QED + +Theorem ODD_SBIT: + ODD (SBIT b n) = (b /\ n = 0) +Proof + rw[SBIT_def, ODD_EXP_IFF] +QED + +(* +Theorem word_to_bytes_from_bin_list: + word_to_bytes (word_from_bin_list ls :'a word) F = + MAP word_from_bin_list (chunks 8 (PAD_RIGHT 0 (dimindex(:'a)) ls)) +Proof + cheat +QED +*) + Datatype: state_array = <| w: num @@ -297,7 +403,7 @@ Proof rw[chi_def] QED -Definition rc_step_def: +Definition rc_step_def[nocompute]: rc_step r = let ra = F :: r in let re = @@ -311,6 +417,9 @@ Definition rc_step_def: TAKE 8 re End +Theorem rc_step_inlined[compute] = + “rc_step r” |> SIMP_CONV (srw_ss()) [rc_step_def, LET_THM]; + Definition rc_def: rc t = if t MOD 255 = 0 then T else @@ -428,6 +537,31 @@ Proof metis_tac[] QED +Theorem LENGTH_pad10s1_less: + m + 2 ≤ x ==> LENGTH (pad10s1 x m) = x - m +Proof + rw[pad10s1_def, LEFT_ADD_DISTRIB, ADD1] + \\ Cases_on`x` \\ gs[] + \\ Cases_on`n` \\ gs[ADD1, LEFT_ADD_DISTRIB] + \\ qmatch_goalsub_rename_tac`2 * k` + \\ `m + (2 * k + (m * k + 2)) = (m + 1) * (k + 2) + (k - m)` by gs[] + \\ pop_assum SUBST_ALL_TAC + \\ DEP_REWRITE_TAC[MOD_TIMES] + \\ simp[] +QED + +Theorem LENGTH_pad10s1_equal: + 2 < x ==> LENGTH (pad10s1 x x) = x +Proof + rw[pad10s1_def,ADD1,LEFT_ADD_DISTRIB] + \\ qmatch_goalsub_abbrev_tac`a MOD x` + \\ `a = (x - 2) + (x * x)` by simp[Abbr`a`] + \\ pop_assum SUBST1_TAC + \\ DEP_ONCE_REWRITE_TAC[GSYM MOD_PLUS] + \\ DEP_REWRITE_TAC[MOD_EQ_0] + \\ simp[] +QED + Definition Keccak_def: Keccak c = sponge (Keccak_p 24) 1600 pad10s1 (1600 - c) End @@ -1791,10 +1925,6 @@ Theorem index_to_triple_100 = EVAL``index_to_triple 4 ^(numLib.term_of_int i)``) |> LIST_CONJ -Definition bool_to_bit_def: - bool_to_bit b = if b then 1 else 0 -End - Definition bools_to_word_def: bools_to_word bs = word_from_bin_list (MAP bool_to_bit bs) @@ -1810,7 +1940,1982 @@ Definition bools_to_hex_string_def: ) $ chunks 8 bs End -open cv_transLib cv_stdTheory; +Definition pad10s1_136_w64_def: + pad10s1_136_w64 (zs: word64 list) (m: word8 list) (a: word64 list list) = + let lm = LENGTH m in + if 136 < lm then let + w64s = MAP concat_word_list $ chunks 8 (TAKE 136 m) + in pad10s1_136_w64 zs (DROP 136 m) ((w64s ++ zs) :: a) + else if lm = 136 then + REVERSE $ + (0x01w::(REPLICATE 15 0w)++(0x8000000000000000w::zs)) :: + ((MAP concat_word_list $ chunks 8 m) ++ zs) :: a + else let + n = 136 - lm; + pad = if n = 1 then [0x81w] else + 0x01w::(REPLICATE (n - 2) 0w)++[0x80w]; + w64s = MAP concat_word_list $ chunks 8 $ m ++ pad + in REVERSE $ (w64s ++ zs) :: a +Termination + WF_REL_TAC`measure $ LENGTH o FST o SND` + \\ rw[LENGTH_DROP] +End + +Theorem TAKE_FLAT_bytes[local]: + ∀n (ls:word8 list). 8 * (n + 1) ≤ LENGTH ls ==> + FLAT (MAP (PAD_RIGHT 0 8 o word_to_bin_list) (TAKE 8 (DROP (8 * n) ls))) = + TAKE 64 (DROP (64 * n) (FLAT (MAP (PAD_RIGHT 0 8 o word_to_bin_list) ls))) +Proof + Induct \\ rw[] + >- ( + Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ simp[TAKE_APPEND] + \\ simp[TAKE_LENGTH_TOO_LONG, LENGTH_PAD_RIGHT_0_8_word_to_bin_list] ) + \\ gs[LEFT_ADD_DISTRIB, MULT_SUC] + \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ qmatch_asmsub_rename_tac`LENGTH ls` \\ Cases_on`ls` \\ gs[] + \\ simp[DROP_APPEND, LENGTH_PAD_RIGHT_0_8_word_to_bin_list, + DROP_LENGTH_TOO_LONG] +QED + +Theorem pad10s1_136_w64_thm: + ∀zs bytes acc bools. + bools = MAP ((=) 1) $ FLAT $ MAP (PAD_RIGHT 0 8 o word_to_bin_list) bytes ∧ + zs = REPLICATE (c DIV 64) 0w ∧ c ≠ 0 ∧ divides 64 c + ⇒ + pad10s1_136_w64 zs bytes acc = + REVERSE acc ++ ( + MAP (MAP word_from_bin_list o chunks 64) $ + MAP (λx. MAP bool_to_bit x ++ REPLICATE c 0) + (chunks 1088 (bools ++ pad10s1 1088 (LENGTH bools))) + ) +Proof + recInduct pad10s1_136_w64_ind + \\ rw[] + \\ simp[Once pad10s1_136_w64_def] + \\ IF_CASES_TAC \\ gs[] + >- ( + qmatch_goalsub_abbrev_tac`lhs = _` + \\ qspecl_then[`136`,`m`](SUBST1_TAC o SYM)TAKE_DROP + \\ qmatch_goalsub_abbrev_tac`MAP _ $ m1 ++ m2` + \\ qmatch_goalsub_abbrev_tac`LENGTH bools` + \\ qunabbrev_tac`lhs` + \\ qmatch_goalsub_abbrev_tac`LENGTH bools2` + \\ qmatch_asmsub_abbrev_tac`bools2 = FLAT (MAP fb m2)` + \\ `bools = FLAT (MAP fb m1) ++ bools2` + by simp[Abbr`bools`] + \\ qunabbrev_tac`bools` + \\ pop_assum SUBST_ALL_TAC + \\ qmatch_goalsub_abbrev_tac`bools1 ++ bools2` + \\ `LENGTH m1 = 136` by simp[Abbr`m1`] + \\ `LENGTH bools1 = 1088` + by ( + simp[Abbr`bools1`, LENGTH_FLAT, MAP_MAP_o, Abbr`fb`] + \\ qmatch_goalsub_abbrev_tac`SUM ls = _` + \\ `ls = REPLICATE 136 8` + by ( + simp[Abbr`ls`, LIST_EQ_REWRITE, EL_MAP, EL_REPLICATE] + \\ simp[PAD_RIGHT] + \\ rpt strip_tac + \\ rewrite_tac[Once ADD_SYM] + \\ irule SUB_ADD + \\ irule LESS_EQ_TRANS + \\ qexists_tac`dimindex(:8)` + \\ simp[LENGTH_word_to_bin_list_bound] ) + \\ simp[SUM_REPLICATE] ) + \\ simp[] + \\ rewrite_tac[GSYM APPEND_ASSOC] + \\ qmatch_goalsub_abbrev_tac`lhs = _` + \\ qmatch_goalsub_abbrev_tac`chunks _ (lm1 ++ lm2)` + \\ qspecl_then[`1088`,`lm1`,`lm2`]mp_tac chunks_append_divides + \\ impl_tac + >- ( + simp[Abbr`lm1`, NULL_EQ, Abbr`lm2`] + \\ conj_tac >- (strip_tac \\ fs[]) + \\ gs[Abbr`bools2`, pad10s1_def] ) + \\ disch_then SUBST_ALL_TAC + \\ `LENGTH lm1 = 1088` by simp[Abbr`lm1`] + \\ first_assum(SUBST1_TAC o SYM) + \\ rewrite_tac[chunks_length] + \\ simp[Abbr`lhs`] + \\ `MAP bool_to_bit lm1 = bools1` + by ( + simp[Abbr`lm1`, MAP_MAP_o, o_DEF] + \\ `bools1 = MAP I bools1` by simp[] + \\ pop_assum SUBST1_TAC + \\ rewrite_tac[MAP_EQ_f, MAP_MAP_o] + \\ simp[Abbr`bools1`, MEM_FLAT, PULL_EXISTS, MEM_MAP, Abbr`fb`, PAD_RIGHT, + MEM_GENLIST, word_to_bin_list_def] + \\ rw[word_to_bin_list_def, bool_to_bit_def] + \\ first_assum(mp_then Any mp_tac MEM_w2l_less) + \\ simp[] ) + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`bools1 ++ r0` + \\ qspecl_then[`64`,`bools1`,`r0`]mp_tac chunks_append_divides + \\ impl_tac + >- ( simp[NULL_EQ, divides_def, Abbr`r0`] \\ strip_tac \\ gs[] ) + \\ disch_then SUBST_ALL_TAC + \\ simp[Abbr`r0`] + \\ qmatch_goalsub_abbrev_tac`l1 ++ l0 = b1 ++ b0` + \\ `LENGTH l0 = c DIV 64` by simp[LENGTH_REPLICATE, Abbr`l0`] + \\ `LENGTH b0 = c DIV 64` by ( + simp[Abbr`b0`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[LENGTH_REPLICATE, NULL_EQ, bool_to_bit_def] ) + \\ `LENGTH l1 = 17` + by ( + simp[Abbr`l1`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[divides_def, bool_to_bit_def, NULL_EQ] + \\ strip_tac \\ gs[] ) + \\ `LENGTH b1 = 17` + by ( + simp[Abbr`b1`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[bool_to_bit_def, divides_def, NULL_EQ] + \\ strip_tac \\ gs[] ) + \\ simp[APPEND_LENGTH_EQ] + \\ conj_tac + >- ( + simp[Abbr`l1`, Abbr`b1`, Abbr`bools1`] + \\ reverse conj_tac + >- ( + simp[Abbr`l0`, Abbr`b0`] + \\ simp[LIST_EQ_REWRITE, EL_REPLICATE] + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[LENGTH_REPLICATE, NULL_EQ] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ simp[LENGTH_REPLICATE, NULL_EQ] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[LENGTH_REPLICATE, NULL_EQ] + \\ simp[REPLICATE_GENLIST, TAKE_GENLIST] + \\ simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST] ) + \\ simp[LIST_EQ_REWRITE] + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ] + \\ (conj_tac \\ strip_tac \\ gs[])) + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ conj_tac >- (conj_tac \\ strip_tac \\ gs[]) + \\ qmatch_goalsub_abbrev_tac`_ ls = _` + \\ `LENGTH ls = 8` by simp[Abbr`ls`] + \\ drule concat_word_list_bytes_to_64 + \\ disch_then SUBST_ALL_TAC + \\ AP_TERM_TAC + \\ simp[Abbr`ls`] + \\ qunabbrev_tac`fb` + \\ irule TAKE_FLAT_bytes + \\ simp[] ) + \\ simp[Abbr`lm2`] + \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ simp[pad10s1_def] + \\ AP_THM_TAC + \\ AP_TERM_TAC + \\ qmatch_goalsub_abbrev_tac`A MOD N = B MOD N` + \\ `B = A + (N - 1) * N` + by ( + qunabbrev_tac`A` + \\ qunabbrev_tac`B` + \\ qunabbrev_tac`N` + \\ simp[LEFT_ADD_DISTRIB] ) + \\ qunabbrev_tac`B` + \\ pop_assum SUBST_ALL_TAC + \\ irule EQ_SYM + \\ ONCE_REWRITE_TAC[ADD_COMM] + \\ irule MOD_TIMES + \\ simp[Abbr`N`] ) + \\ qmatch_goalsub_abbrev_tac`MAP _ $ MAP _ $ chunks _ (l1 ++ l2)` + \\ `LENGTH l1 = 8 * LENGTH m` + by ( + simp[Abbr`l1`, LENGTH_FLAT, MAP_MAP_o] + \\ qmatch_goalsub_abbrev_tac`SUM ls` + \\ `ls = REPLICATE (LENGTH m) 8` + by simp[Abbr`ls`, LIST_EQ_REWRITE, EL_MAP, EL_REPLICATE] + \\ simp[SUM_REPLICATE] ) + \\ IF_CASES_TAC \\ gs[] + >- ( + DEP_REWRITE_TAC[chunks_append_divides] + \\ simp[NULL_EQ] + \\ conj_tac + >- ( + conj_tac >- ( strip_tac \\ gs[] ) + \\ gs[Abbr`l1`,Abbr`l2`,pad10s1_def] ) + \\ qpat_assum`LENGTH l1 = _`(SUBST1_TAC o SYM) + \\ simp[chunks_length, PULL_EXISTS] + \\ `LENGTH l2 = 1088` + by ( simp[Abbr`l2`] \\ gs[Abbr`l1`] \\ simp[pad10s1_def] ) + \\ pop_assum(SUBST1_TAC o SYM) + \\ simp[chunks_length] + \\ conj_tac + >- ( + DEP_ONCE_REWRITE_TAC[chunks_append_divides] + \\ gs[NULL_EQ] + \\ conj_tac + >- ( conj_tac >- rw[divides_def] + \\ strip_tac \\ gs[Abbr`l1`]) + \\ qmatch_goalsub_abbrev_tac`m1 ++ m2 = m3 ++ m4` + \\ `LENGTH m1 = 17` + by ( + simp[Abbr`m1`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[divides_def, NULL_EQ, bool_to_bit_def] + \\ strip_tac \\ gs[] ) + \\ `LENGTH m3 = 17` + by ( + simp[Abbr`m3`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ gs[bool_to_bit_def, NULL_EQ, divides_def] + \\ strip_tac \\ gs[Abbr`l1`] + \\ strip_tac \\ gs[] ) + \\ `LENGTH m2 = c DIV 64` by simp[Abbr`m2`] + \\ `LENGTH m4 = c DIV 64` by ( + simp[Abbr`m4`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[bool_to_bit_def, NULL_EQ] ) + \\ simp[APPEND_LENGTH_EQ] + \\ conj_tac + >- ( + gs[Abbr`m1`,Abbr`m3`,LIST_EQ_REWRITE,EL_MAP] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ conj_tac + >- ( + conj_tac >- (strip_tac \\ gs[]) + \\ strip_tac \\ gs[Abbr`l1`]) + \\ gs[Abbr`l1`] + \\ simp[MAP_MAP_o, MAP_FLAT] + \\ qmatch_goalsub_abbrev_tac`MAP (f o g)` + \\ `MAP (f o g) m = MAP g m` + by ( + rw[Abbr`f`,Abbr`g`, MAP_EQ_f] + \\ qmatch_goalsub_abbrev_tac`MAP _ l` + \\ rw[LIST_EQ_REWRITE, EL_MAP] + \\ rw[bool_to_bit_def] + \\ qmatch_goalsub_abbrev_tac`b =0` + \\ `MEM b l` by metis_tac[MEM_EL] + \\ pop_assum mp_tac + \\ simp[Abbr`l`, PAD_RIGHT, word_to_bin_list_def, MEM_GENLIST] + \\ rw[] + \\ `b < 2` suffices_by rw[] + \\ qspec_then`e`irule(Q.GEN`w`MEM_w2l_less) + \\ simp[] + \\ metis_tac[]) + \\ pop_assum SUBST_ALL_TAC + \\ qunabbrev_tac`g` + \\ DEP_REWRITE_TAC[concat_word_list_bytes_to_64] + \\ simp[] + \\ DEP_REWRITE_TAC[TAKE_FLAT_bytes] + \\ simp[] ) + \\ gs[Abbr`m2`, Abbr`m4`, LIST_EQ_REWRITE, EL_MAP, EL_REPLICATE] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST, REPLICATE_GENLIST, TAKE_GENLIST]) + \\ gs[Abbr`l1`] + \\ simp[Abbr`l2`] + \\ DEP_REWRITE_TAC[chunks_append_divides] + \\ simp[NULL_EQ, LENGTH_pad10s1_equal] + \\ conj_tac >- rw[divides_def, pad10s1_def] + \\ once_rewrite_tac[CONS_APPEND] + \\ rewrite_tac[APPEND_ASSOC] + \\ qmatch_goalsub_abbrev_tac`m1 ++ m2 = m3 ++ m4` + \\ `LENGTH m2 = c DIV 64` by simp[Abbr`m2`] + \\ `LENGTH m4 = c DIV 64` by ( + simp[Abbr`m4`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[bool_to_bit_def, NULL_EQ] ) + \\ `LENGTH m1 = 17` by simp[Abbr`m1`] + \\ `LENGTH m3 = 17` by ( + simp[Abbr`m3`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[bool_to_bit_def, LENGTH_pad10s1_equal, NULL_EQ, divides_def] + \\ simp[pad10s1_def] ) + \\ simp[APPEND_LENGTH_EQ] + \\ reverse conj_tac + >- ( + gs[Abbr`m2`, Abbr`m4`, LIST_EQ_REWRITE, EL_MAP, EL_REPLICATE] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST, REPLICATE_GENLIST, TAKE_GENLIST] ) + \\ gs[Abbr`m1`,Abbr`m3`] + \\ simp[pad10s1_def, bool_to_bit_def] + \\ simp[Once chunks_def] + \\ conj_tac + >- ( + simp[TAKE_APPEND] + \\ rewrite_tac[REPLICATE_GENLIST, TAKE_GENLIST] + \\ simp[] ) + \\ rewrite_tac[DROP_APPEND, REPLICATE_GENLIST, DROP_GENLIST, LENGTH_GENLIST] + \\ qmatch_goalsub_abbrev_tac`chunks _ (GENLIST _ n ++ _)` + \\ simp[] + \\ simp[Once chunks_def] + \\ `~(n + 1 <= 64)` by simp[Abbr`n`] + \\ `64 - n = 0` by simp[Abbr`n`] + \\ simp[TAKE_APPEND] + \\ conj_tac + >- ( + simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST, REPLICATE_GENLIST, TAKE_GENLIST] ) + \\ simp[DROP_APPEND, DROP_GENLIST] + \\ qunabbrev_tac`n` + \\ rpt ( + qmatch_goalsub_abbrev_tac`chunks _ (GENLIST _ n ++ _)` + \\ simp[Once chunks_def] + \\ `~(n + 1 <= 64)` by simp[Abbr`n`] + \\ `64 - n = 0` by simp[Abbr`n`] + \\ simp[TAKE_APPEND] + \\ conj_tac + >- ( + simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST, REPLICATE_GENLIST, TAKE_GENLIST] ) + \\ simp[DROP_APPEND, DROP_GENLIST] + \\ qunabbrev_tac`n`) + \\ gs[] + \\ simp[Once chunks_def]) + \\ simp[PULL_EXISTS] + \\ `LENGTH (l1 ++ l2) = 1088` + by ( + gs[Abbr`l2`, Abbr`l1`] + \\ DEP_REWRITE_TAC[LENGTH_pad10s1_less] + \\ gs[] ) + \\ simp[Once chunks_def] + \\ qmatch_goalsub_abbrev_tac`chunks 8 ls` + \\ `LENGTH ls = 136` by rw[Abbr`ls`] + \\ DEP_ONCE_REWRITE_TAC[chunks_append_divides] + \\ gs[NULL_EQ] + \\ conj_tac + >- ( + conj_tac >- rw[divides_def] + \\ strip_tac \\ gs[Abbr`l1`] + \\ strip_tac \\ gs[] ) + \\ qmatch_goalsub_abbrev_tac`m1 ++ m2 = m3 ++ m4` + \\ `LENGTH m1 = 17` + by ( + simp[Abbr`m1`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[divides_def, NULL_EQ, bool_to_bit_def] + \\ strip_tac \\ gs[] ) + \\ `LENGTH m3 = 17` + by ( + simp[Abbr`m3`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ gs[bool_to_bit_def, NULL_EQ, divides_def] + \\ strip_tac \\ gs[Abbr`l1`] + \\ strip_tac \\ gs[] ) + \\ `LENGTH m2 = c DIV 64` by simp[Abbr`m2`] + \\ `LENGTH m4 = c DIV 64` by ( + simp[Abbr`m4`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[bool_to_bit_def, NULL_EQ] ) + \\ simp[APPEND_LENGTH_EQ] + \\ conj_tac + >- ( + gs[Abbr`m1`,Abbr`m3`,LIST_EQ_REWRITE,EL_MAP] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ conj_tac + >- ( + conj_tac >- (strip_tac \\ gs[]) + \\ strip_tac \\ gs[Abbr`l1`] + \\ strip_tac \\ gs[] ) + \\ DEP_REWRITE_TAC[concat_word_list_bytes_to_64] + \\ simp[] + \\ DEP_REWRITE_TAC[TAKE_FLAT_bytes] + \\ simp[] + \\ AP_TERM_TAC + \\ qunabbrev_tac`ls` + \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ simp[Abbr`l1`] + \\ qmatch_goalsub_abbrev_tac`p1 ++ p2 = p3 ++ p4` + \\ `p1 = p3` + by ( + simp[Abbr`p1`,Abbr`p3`,MAP_MAP_o] + \\ qmatch_goalsub_abbrev_tac`ls = _` + \\ simp[LIST_EQ_REWRITE, EL_MAP, bool_to_bit_def] + \\ rw[] + \\ qmatch_goalsub_abbrev_tac`b = 0` + \\ `MEM b ls` by metis_tac[MEM_EL] + \\ pop_assum mp_tac + \\ simp[Abbr`ls`, PAD_RIGHT, word_to_bin_list_def, + PULL_EXISTS, MEM_GENLIST, MEM_MAP, MEM_FLAT] + \\ qx_gen_tac`e` \\ rw[] + \\ `b < 2` suffices_by rw[] + \\ qspec_then`e`irule(Q.GEN`w`MEM_w2l_less) + \\ simp[] + \\ metis_tac[]) + \\ simp[] + \\ gs[Abbr`p2`,Abbr`p4`] + \\ gs[Abbr`l2`] + \\ IF_CASES_TAC + >- ( + `8 * LENGTH m = 1080` by gs[] + \\ simp[pad10s1_def, bool_to_bit_def, REPLICATE_GENLIST, PAD_RIGHT] ) + \\ simp[pad10s1_def, PAD_RIGHT, bool_to_bit_def] + \\ simp[REPLICATE_GENLIST] + \\ qmatch_goalsub_abbrev_tac`_ = GENLIST _ (n MOD 1088)` + \\ gs[] + \\ gs[pad10s1_def, ADD1] + \\ `n MOD 1088 = 1086 - 8 * LENGTH m` by gs[] + \\ pop_assum SUBST1_TAC + \\ simp[LIST_EQ_REWRITE, ADD1, LENGTH_FLAT] + \\ qmatch_goalsub_abbrev_tac`SUM ls` + \\ `ls = REPLICATE (134 - LENGTH m) 8` + by simp[Abbr`ls`, LIST_EQ_REWRITE, EL_REPLICATE, EL_MAP] + \\ conj_asm1_tac >- simp[SUM_REPLICATE] + \\ rw[] + \\ qmatch_goalsub_abbrev_tac`EL i ls` + \\ `LENGTH ls = 1086 - 8 * LENGTH m` + by simp[Abbr`ls`, ADD1, LENGTH_FLAT] + \\ `MEM (EL i ls) ls` by metis_tac[MEM_EL] + \\ `EVERY (λx. x = 0) ls` suffices_by simp[EVERY_MEM] + \\ simp[Abbr`ls`, EVERY_FLAT, EVERY_GENLIST] ) + \\ gs[Abbr`m2`, Abbr`m4`, LIST_EQ_REWRITE, EL_MAP, EL_REPLICATE] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ] + \\ simp[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by simp[] + \\ qunabbrev_tac`A` + \\ simp[l2n_eq_0, EVERY_GENLIST, REPLICATE_GENLIST, TAKE_GENLIST] +QED + +Definition state_bools_w64_def: + state_bools_w64 bools (w64s:word64 list) = + ((LENGTH bools = 1600) ∧ + (w64s = MAP word_from_bin_list $ chunks 64 $ MAP bool_to_bit bools)) +End + +Definition bytes_to_bools_def: + bytes_to_bools (bytes: word8 list) = + MAP ((=) 1) (FLAT (MAP (PAD_RIGHT 0 8 o word_to_bin_list) bytes)) +End + +Theorem pad10s1_136_w64_sponge_init: + let + N = bytes_to_bools bytes; + r = 1600 - 512; + P = N ++ pad10s1 r (LENGTH N); + c = 1600 - r; + Pis = chunks r P; + bs = MAP (λPi. Pi ++ REPLICATE c F) Pis + in + EVERY2 state_bools_w64 bs + (pad10s1_136_w64 (REPLICATE (c DIV 64) 0w) bytes []) +Proof + rw[] + \\ qabbrev_tac`c = 512` + \\ qmatch_goalsub_abbrev_tac`LENGTH bools` + \\ DEP_REWRITE_TAC[pad10s1_136_w64_thm] + \\ simp[GSYM bytes_to_bools_def] + \\ simp[Abbr`c`, divides_def] + \\ simp[LIST_REL_EL_EQN, EL_MAP] + \\ simp[state_bools_w64_def, bool_to_bit_def] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ conj_asm1_tac + >- ( simp[NULL_EQ] \\ simp[Once pad10s1_def] ) + \\ DEP_REWRITE_TAC[LENGTH_TAKE] + \\ simp[DROP_APPEND] + \\ pop_assum mp_tac + \\ simp[NULL_EQ] \\ strip_tac + \\ qpat_x_assum`n < _`mp_tac + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ] + \\ qspecl_then[`1088`,`LENGTH bools`]mp_tac(Q.GENL[`x`,`m`] LENGTH_pad10s1) + \\ simp[] \\ strip_tac \\ simp[bool_to_bit_def] +QED + +Theorem xor_state_bools_w64: + state_bools_w64 b1 w1 /\ + state_bools_w64 b2 w2 + ==> + state_bools_w64 (MAP2 (λx y. x ≠ y) b1 b2) + (MAP2 word_xor w1 w2) +Proof + rw[state_bools_w64_def] + \\ DEP_REWRITE_TAC[MAP2_MAP] + \\ simp[chunks_MAP] + \\ `LENGTH (chunks 64 b1) = 25 ∧ LENGTH (chunks 64 b2) = 25` + by ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ, bool_to_bit_def, divides_def] + \\ rw[] \\ strip_tac \\ gs[] ) + \\ simp[chunks_ZIP] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 1 |> DISCH_ALL] + \\ simp[] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL] + \\ simp[MAP_MAP_o] + \\ simp[MAP_EQ_f, FORALL_PROD, MEM_ZIP, PULL_EXISTS] + \\ rw[] + \\ qmatch_goalsub_abbrev_tac`bs1,bs2` + \\ `LENGTH bs1 = 64 ∧ LENGTH bs2 = 64` + by ( + simp[Abbr`bs1`,Abbr`bs2`] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ gs[NULL_EQ, bool_to_bit_def, divides_def] + \\ rw[] \\ strip_tac \\ gs[] ) + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ simp[] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 1 |> DISCH_ALL] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL] + \\ simp[MAP_MAP_o] + \\ AP_TERM_TAC + \\ simp[MAP_EQ_f, FORALL_PROD] + \\ rw[bool_to_bit_def] +QED + +Definition theta_c_w64_def: + theta_c_w64 (s:word64 list) = [ + EL 0 s ?? EL 5 s ?? EL 10 s ?? EL 15 s ?? EL 20 s; + EL 1 s ?? EL 6 s ?? EL 11 s ?? EL 16 s ?? EL 21 s; + EL 2 s ?? EL 7 s ?? EL 12 s ?? EL 17 s ?? EL 22 s; + EL 3 s ?? EL 8 s ?? EL 13 s ?? EL 18 s ?? EL 23 s; + EL 4 s ?? EL 9 s ?? EL 14 s ?? EL 19 s ?? EL 24 s + ] +End + +Theorem EL_theta_c_w64: + i < 5 ==> + EL i (theta_c_w64 s) = + EL (i ) s ?? + EL (i + 5) s ?? + EL (i + 10) s ?? + EL (i + 15) s ?? + EL (i + 20) s +Proof + rw[theta_c_w64_def, NUMERAL_LESS_THM] + \\ rw[] +QED + +Theorem theta_c_w64_thm: + state_bools_w64 bs ws /\ + string_to_state_array bs = s /\ + i < 5 + ==> + EL i (theta_c_w64 ws) = + word_from_bin_list $ MAP bool_to_bit $ + GENLIST (λj. theta_c s.A (i, j)) 64 +Proof + strip_tac + \\ qmatch_goalsub_abbrev_tac`GENLIST _ n` + \\ gvs[state_bools_w64_def] + \\ rw[string_to_state_array_def, b2w_def, EL_theta_c_w64] + \\ DEP_REWRITE_TAC[EL_MAP] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ, bool_to_bit_def] + \\ conj_tac >- ( rw[Abbr`n`] \\ strip_tac \\ gs[] ) + \\ conj_tac >- rw[Abbr`n`, divides_def] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ qmatch_goalsub_abbrev_tac`GENLIST f` + \\ simp[NULL_EQ, bool_to_bit_def] + \\ conj_tac >- ( rw[Abbr`n`] \\ strip_tac \\ gs[] ) + \\ conj_tac >- ( rw[Abbr`n`, divides_def] \\ strip_tac \\ gs[] ) + \\ simp[GSYM MAP_DROP, GSYM MAP_TAKE] + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ conj_tac >- simp[Abbr`n`] + \\ AP_TERM_TAC + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 1 |> DISCH_ALL] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL] + \\ rpt (conj_tac >- simp[Abbr`n`]) + \\ simp[MAP_MAP_o] + \\ simp[o_DEF] + \\ simp[LIST_EQ_REWRITE] + \\ conj_tac >- simp[Abbr`n`] + \\ gen_tac + \\ DEP_REWRITE_TAC[LENGTH_TAKE] + \\ conj_tac >- simp[Abbr`n`] + \\ strip_tac + \\ simp[EL_MAP] + \\ DEP_REWRITE_TAC[EL_MAP] + \\ conj_tac >- simp[Abbr`n`] + \\ DEP_REWRITE_TAC[EL_ZIP, EL_TAKE, EL_DROP] + \\ conj_tac >- simp[Abbr`n`] + \\ simp[] + \\ simp[Abbr`f`] + \\ simp[theta_c_def, restrict_def] + \\ rw[bool_to_bit_def] +QED + +Definition theta_d_w64_def: + theta_d_w64 (s:word64 list) = + let c = theta_c_w64 s in + GENLIST (λx. + let + idx1 = (x + 4) MOD 5; + idx0 = (x + 1) MOD 5; + b0 = EL idx0 c; + in word_rol b0 1 ?? EL idx1 c + ) 5 +End + +Theorem theta_d_w64_thm: + state_bools_w64 bs ws /\ + string_to_state_array bs = s /\ + i < 5 + ==> + EL i (theta_d_w64 ws) = + word_from_bin_list $ MAP bool_to_bit $ + GENLIST (λj. theta_d s.w s.A (i,j)) 64 +Proof + strip_tac + \\ qmatch_goalsub_abbrev_tac`GENLIST _ n` + \\ simp[theta_d_def] + \\ rewrite_tac[theta_d_w64_def] + \\ BasicProvers.LET_ELIM_TAC + \\ qmatch_goalsub_abbrev_tac`GENLIST _ m` + \\ DEP_REWRITE_TAC[EL_GENLIST] + \\ conj_tac >- simp[] + \\ BasicProvers.LET_ELIM_TAC + \\ qunabbrev_tac`c` + \\ qunabbrev_tac`b0` + \\ DEP_REWRITE_TAC[theta_c_w64_thm] + \\ conj_tac >- gs[Abbr`idx0`,Abbr`idx1`, Abbr`m`] + \\ asm_simp_tac std_ss [] + \\ DEP_REWRITE_TAC[word_from_bin_list_rol] + \\ simp[] + \\ conj_asm1_tac >- rw[Abbr`n`] + \\ `MIN (n - 1) n = n - 1` by simp[Abbr`n`] + \\ simp[LASTN_DROP, BUTLASTN_TAKE] + \\ simp[GSYM MAP_DROP, GSYM MAP_TAKE, DROP_GENLIST, TAKE_GENLIST] + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ conj_tac >- simp[Abbr`n`] + \\ AP_TERM_TAC + \\ rewrite_tac[GSYM MAP_APPEND, GSYM MAP] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 1 |> DISCH_ALL] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL] + \\ rpt (conj_tac >- simp[Abbr`n`]) + \\ simp[MAP_MAP_o, o_DEF] + \\ `s.w = n` by ( rw[string_to_state_array_def] \\ gs[state_bools_w64_def, b2w_def] ) + \\ simp[LIST_EQ_REWRITE, EL_APPEND_EQN, ADD1, EL_MAP, EL_ZIP, PRE_SUB1] + \\ Cases \\ simp[bool_to_bit_def] >- rw[] + \\ simp[ADD1] \\ rw[] +QED + +Theorem LENGTH_theta_d_w64[simp]: + LENGTH (theta_d_w64 s) = 5 +Proof + rw[theta_d_w64_def] +QED + +Definition theta_w64_def: + theta_w64 s = + let t = theta_d_w64 s in + GENLIST (λi. EL i s ?? EL (i MOD 5) t) 25 +End + +Theorem LENGTH_theta_w64[simp]: + LENGTH (theta_w64 s) = 25 +Proof + rw[theta_w64_def] +QED + +Theorem theta_w64_thm: + state_bools_w64 bs ws /\ + string_to_state_array bs = s + ==> + state_bools_w64 (state_array_to_string (theta s)) + $ theta_w64 ws +Proof + simp[state_bools_w64_def] + \\ strip_tac + \\ conj_tac + >- gvs[string_to_state_array_def, b2w_def] + \\ simp[theta_def] + \\ simp[LIST_EQ_REWRITE] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ] + \\ gvs[string_to_state_array_def, b2w_def, divides_def, bool_to_bit_def] + \\ rewrite_tac[state_array_to_string_def] + \\ disch_then(mp_tac o Q.AP_TERM`LENGTH`) + \\ simp[] ) + \\ simp[EL_MAP] + \\ rewrite_tac[theta_w64_def] + \\ gen_tac \\ strip_tac + \\ BasicProvers.LET_ELIM_TAC + \\ DEP_REWRITE_TAC[EL_GENLIST, EL_MAP] + \\ conj_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ gs[NULL_EQ] + \\ strip_tac \\ gs[] ) + \\ qunabbrev_tac`t` + \\ DEP_REWRITE_TAC[theta_d_w64_thm] + \\ conj_tac + >- simp[state_bools_w64_def, DIV_LT_X] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ conj_tac + >- ( + gs[NULL_EQ] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ gs[NULL_EQ] + \\ rpt strip_tac \\ gs[] ) + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ conj_tac >- simp[] + \\ AP_TERM_TAC + \\ rewrite_tac[GSYM MAP_DROP, GSYM MAP_TAKE] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 1 |> DISCH_ALL] + \\ DEP_REWRITE_TAC[ZIP_MAP |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL] + \\ conj_tac >- rw[] + \\ conj_tac >- rw[] + \\ rewrite_tac[MAP_MAP_o] + \\ qmatch_goalsub_abbrev_tac`ZIP ls` + \\ `s.w = 64` + by ( rw[] \\ simp[string_to_state_array_def, b2w_def]) + \\ simp[MAP_MAP_o, LIST_EQ_REWRITE, EL_MAP] + \\ conj_asm1_tac + >- ( simp[Abbr`ls`, LENGTH_TAKE_EQ]) + \\ simp[EL_MAP, EL_TAKE, EL_ZIP, Abbr`ls`] + \\ qx_gen_tac`i` \\ strip_tac + \\ DEP_REWRITE_TAC[EL_DROP] + \\ simp[] + \\ rewrite_tac[state_array_to_string_compute] + \\ DEP_REWRITE_TAC[EL_GENLIST] + \\ simp[] + \\ simp[restrict_def, DIV_LT_X] + \\ `i DIV 64 = 0` by simp[DIV_EQ_0] + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`EL ix bs` + \\ qmatch_goalsub_abbrev_tac`s.A t` + \\ `EL ix bs = s.A t` suffices_by rw[bool_to_bit_def] + \\ rw[string_to_state_array_def, b2w_def, Abbr`t`, restrict_def, DIV_LT_X] + \\ rw[Abbr`ix`] + \\ AP_THM_TAC + \\ AP_TERM_TAC + \\ simp[] + \\ qspec_then`5`mp_tac DIVISION + \\ impl_tac >- rw[] + \\ disch_then(qspec_then`x`mp_tac) + \\ simp[] +QED + +Theorem rho_xy_invert: + t <= 23 ==> + (rho_xy t = (1,0) ==> t = 00) ∧ + (rho_xy t = (2,0) ==> t = 18) ∧ + (rho_xy t = (3,0) ==> t = 06) ∧ + (rho_xy t = (4,0) ==> t = 12) ∧ + (rho_xy t = (0,1) ==> t = 07) ∧ + (rho_xy t = (1,1) ==> t = 23) ∧ + (rho_xy t = (2,1) ==> t = 02) ∧ + (rho_xy t = (3,1) ==> t = 09) ∧ + (rho_xy t = (4,1) ==> t = 22) ∧ + (rho_xy t = (0,2) ==> t = 01) ∧ + (rho_xy t = (1,2) ==> t = 03) ∧ + (rho_xy t = (2,2) ==> t = 17) ∧ + (rho_xy t = (3,2) ==> t = 16) ∧ + (rho_xy t = (4,2) ==> t = 20) ∧ + (rho_xy t = (0,3) ==> t = 13) ∧ + (rho_xy t = (1,3) ==> t = 08) ∧ + (rho_xy t = (2,3) ==> t = 04) ∧ + (rho_xy t = (3,3) ==> t = 05) ∧ + (rho_xy t = (4,3) ==> t = 15) ∧ + (rho_xy t = (0,4) ==> t = 19) ∧ + (rho_xy t = (1,4) ==> t = 10) ∧ + (rho_xy t = (2,4) ==> t = 21) ∧ + (rho_xy t = (3,4) ==> t = 14) ∧ + (rho_xy t = (4,4) ==> t = 11) +Proof + simp[LESS_OR_EQ] + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ EVAL_TAC +QED + +Definition rho_w64_shifts_def: + rho_w64_shifts = + MAP (λt. (19200 - (((t + 1) * (t + 2)) DIV 2)) MOD 64) + [00 ;18 ;06 ;12 ;07 ;23 ;02 ;09 ;22 ;01 ;03 ;17 ;16 + ;20 ;13 ;08 ;04 ;05 ;15 ;19 ;10 ;21 ;14 ;11] +End + +Theorem LENGTH_rho_w64_shifts[simp]: + LENGTH rho_w64_shifts = 24 +Proof + rw[rho_w64_shifts_def] +QED + +Definition rho_w64_def: + rho_w64 (s: word64 list) = + HD s :: + GENLIST (λi. + word_ror (EL (SUC i) s) (EL i rho_w64_shifts) + ) 24 +End + +Theorem rho_w64_thm: + state_bools_w64 bs ws ∧ + string_to_state_array bs = s ⇒ + state_bools_w64 (state_array_to_string (rho s)) (rho_w64 ws) +Proof + simp[state_bools_w64_def] + \\ strip_tac + \\ conj_asm1_tac + >- rw[string_to_state_array_def, b2w_def] + \\ rewrite_tac[rho_def, rho_w64_def] + \\ qmatch_goalsub_abbrev_tac`GENLIST _ n` + \\ rewrite_tac[state_array_to_string_compute] + \\ qmatch_goalsub_abbrev_tac`5 * 5 * sw` + \\ `sw = s.w` by rw[Abbr`sw`] + \\ qabbrev_tac`m = 5 * 5 * sw` + \\ simp[restrict_def] + \\ qmatch_goalsub_abbrev_tac`GENLIST f m` + \\ simp[LIST_EQ_REWRITE] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ rewrite_tac[NULL_LENGTH, LENGTH_MAP, LENGTH_GENLIST] + \\ conj_asm1_tac >- rw[Abbr`m`] + \\ conj_asm1_tac >- rw[Abbr`m`, divides_def, bool_to_bit_def, Abbr`n`] + \\ pop_assum(assume_tac o SYM) + \\ Cases + >- ( + simp[] + \\ rewrite_tac[GSYM EL] + \\ DEP_REWRITE_TAC[EL_MAP] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ rewrite_tac[NULL_LENGTH, LENGTH_MAP, LENGTH_GENLIST] + \\ simp[] \\ conj_tac >- (strip_tac \\ gs[]) + \\ AP_TERM_TAC + \\ rewrite_tac[GSYM EL] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ rewrite_tac[NULL_LENGTH, LENGTH_MAP, LENGTH_GENLIST] + \\ simp[] \\ conj_tac >- (strip_tac \\ gs[]) + \\ conj_tac >- (strip_tac \\ gs[]) + \\ simp[LIST_EQ_REWRITE, LENGTH_TAKE_EQ] + \\ conj_tac >- rw[Abbr`m`] + \\ rw[] + \\ DEP_REWRITE_TAC[EL_TAKE, EL_MAP, EL_GENLIST] + \\ conj_tac >- rw[Abbr`m`] + \\ simp[Abbr`f`, DIV_EQ_0, DIV_LT_X] + \\ rw[string_to_state_array_def, restrict_def, b2w_def] ) + \\ simp[] + \\ strip_tac + \\ DEP_REWRITE_TAC[EL_MAP] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ rewrite_tac[NULL_LENGTH, LENGTH_MAP, LENGTH_GENLIST] + \\ simp[] \\ conj_tac >- (strip_tac \\ gs[]) + \\ conj_tac >- ( rw[bool_to_bit_def, divides_def] \\ gvs[Abbr`n`] ) + \\ DEP_REWRITE_TAC[EL_chunks] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ rewrite_tac[NULL_LENGTH, LENGTH_MAP, LENGTH_GENLIST] + \\ simp[] \\ conj_tac >- (strip_tac \\ gs[]) + \\ conj_tac >- ( rw[bool_to_bit_def, divides_def] \\ gvs[Abbr`n`] ) + \\ simp[GSYM MAP_DROP, GSYM TAKE_DROP, DROP_GENLIST] + \\ DEP_REWRITE_TAC[word_from_bin_list_ror] + \\ simp[LENGTH_TAKE_EQ] + \\ qmatch_asmsub_rename_tac`p < n` + \\ `p < 24` by gs[Abbr`n`] + \\ simp[] + \\ conj_asm1_tac + >- ( + rewrite_tac[rho_w64_shifts_def] + \\ DEP_REWRITE_TAC[EL_MAP] + \\ simp[] ) + \\ AP_TERM_TAC + \\ simp[LIST_EQ_REWRITE] + \\ DEP_REWRITE_TAC[LENGTH_TAKE_EQ, LENGTH_BUTLASTN, LENGTH_LASTN] + \\ simp[] + \\ conj_tac >- rw[Abbr`m`] + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_APPEND_EQN] + \\ DEP_REWRITE_TAC[LENGTH_TAKE_EQ, LENGTH_BUTLASTN, LENGTH_LASTN] + \\ simp[] + \\ DEP_REWRITE_TAC[EL_TAKE, EL_MAP, EL_GENLIST] + \\ simp[] + \\ conj_tac >- ( + rw[Abbr`m`] + \\ qpat_x_assum`p < _`mp_tac + \\ simp_tac std_ss [NUMERAL_LESS_THM] + \\ strip_tac \\ BasicProvers.VAR_EQ_TAC \\ EVAL_TAC ) + \\ DEP_REWRITE_TAC[LASTN_DROP, BUTLASTN_TAKE] + \\ simp[] + \\ `x DIV 64 = 0` by simp[DIV_EQ_0] + \\ qunabbrev_tac`f` + \\ simp[DIV_LT_X] + \\ qmatch_goalsub_abbrev_tac`t + 2` + \\ IF_CASES_TAC + >- ( + DEP_REWRITE_TAC[EL_DROP, EL_TAKE, EL_MAP] + \\ simp[] + \\ AP_TERM_TAC + \\ IF_CASES_TAC + >- ( + pop_assum mp_tac + \\ qpat_x_assum`p < 24`mp_tac + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac \\ simp[] + \\ simp[Abbr`n`] ) + \\ qunabbrev_tac`t` + \\ numLib.LEAST_ELIM_TAC + \\ first_assum(mp_then Any mp_tac rho_xy_exists) + \\ impl_tac >- simp[DIV_LT_X] + \\ strip_tac + \\ conj_tac >- (qexists_tac`t` \\ simp[]) + \\ qx_gen_tac`t2` + \\ strip_tac + \\ `¬(t < t2)` by metis_tac[] + \\ `t2 <= 23` by gs[] + \\ rw[string_to_state_array_def, restrict_def, DIV_LT_X, b2w_def] + \\ AP_THM_TAC + \\ AP_TERM_TAC + \\ qpat_x_assum`_ = _.w`kall_tac + \\ qpat_x_assum`x < _ -_`mp_tac + \\ rewrite_tac[rho_w64_shifts_def] + \\ qunabbrev_tac`n` + \\ drule rho_xy_invert \\ strip_tac + \\ qpat_x_assum`x < _`mp_tac + \\ qpat_x_assum`p < _`mp_tac + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ qpat_x_assum`rho_xy _ = _`mp_tac + \\ simp_tac std_ss [] \\ strip_tac + \\ first_x_assum drule + \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ DEP_REWRITE_TAC[EL_MAP] + \\ simp_tac (srw_ss()) [] + \\ rw[]) + \\ DEP_REWRITE_TAC[EL_DROP] + \\ simp[] + \\ AP_TERM_TAC + \\ IF_CASES_TAC + >- ( + pop_assum mp_tac + \\ qpat_x_assum`p < 24`mp_tac + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac \\ simp[] + \\ simp[Abbr`n`] ) + \\ qunabbrev_tac`t` + \\ numLib.LEAST_ELIM_TAC + \\ first_assum(mp_then Any mp_tac rho_xy_exists) + \\ impl_tac >- simp[DIV_LT_X] + \\ strip_tac + \\ conj_tac >- (qexists_tac`t` \\ simp[]) + \\ qx_gen_tac`t2` + \\ strip_tac + \\ `¬(t < t2)` by metis_tac[] + \\ `t2 <= 23` by gs[] + \\ rw[string_to_state_array_def, restrict_def, DIV_LT_X, b2w_def] + \\ AP_THM_TAC + \\ AP_TERM_TAC + \\ qpat_x_assum`_ = _.w`kall_tac + \\ qpat_x_assum`~(x < _ -_)`mp_tac + \\ rewrite_tac[rho_w64_shifts_def] + \\ qunabbrev_tac`n` + \\ drule rho_xy_invert \\ strip_tac + \\ qpat_x_assum`x < _`mp_tac + \\ qpat_x_assum`p < _`mp_tac + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ qpat_x_assum`rho_xy _ = _`mp_tac + \\ simp_tac std_ss [] \\ strip_tac + \\ first_x_assum drule + \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ DEP_REWRITE_TAC[EL_MAP] + \\ simp_tac (srw_ss()) [] + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac \\ BasicProvers.VAR_EQ_TAC + \\ EVAL_TAC +QED + +Theorem rho_w64_MAP2: + LENGTH s = 25 ==> + rho_w64 s = + case s of h::t => + h :: MAP2 (λx y. word_ror x y) t rho_w64_shifts + | _ => [] +Proof + strip_tac + \\ CASE_TAC >- fs[] + \\ rewrite_tac[LIST_EQ_REWRITE] + \\ conj_tac >- gs[rho_w64_def, LENGTH_TL, ADD1, MIN_DEF] + \\ Cases >- rw[rho_w64_def] + \\ strip_tac + \\ rewrite_tac[rho_w64_def] + \\ DEP_REWRITE_TAC[EL_CONS, EL_GENLIST, EL_MAP2] + \\ conj_tac >- fs[rho_w64_def, LENGTH_TL] + \\ simp[] +QED + +Definition pi_w64_indices_def: + pi_w64_indices = + GENLIST (λi. let + x = i MOD 5; + y = i DIV 5; + x' = (x + 3 * y) MOD 5 in + x' + 5 * x) 25 +End + +Theorem pi_w64_indices_eq = EVAL ``pi_w64_indices``; + +Theorem pi_w64_indices_bound: + EVERY ($> 25) pi_w64_indices +Proof + rw[pi_w64_indices_eq] +QED + +Definition pi_w64_def: + pi_w64 (s: word64 list) = + MAP (λi. EL i s) pi_w64_indices +End + +Theorem pi_w64_thm: + state_bools_w64 bs ws ∧ + string_to_state_array bs = s ⇒ + state_bools_w64 (state_array_to_string (pi s)) (pi_w64 ws) +Proof + strip_tac + \\ gs[state_bools_w64_def] + \\ conj_asm1_tac >- rw[string_to_state_array_def, b2w_def] + \\ simp[pi_def, pi_w64_def] + \\ rewrite_tac[state_array_to_string_compute] + \\ qmatch_goalsub_abbrev_tac`5 * 5 * sw` + \\ `sw = s.w` by rw[Abbr`sw`] + \\ qunabbrev_tac`sw` \\ pop_assum SUBST_ALL_TAC + \\ qmatch_goalsub_abbrev_tac`GENLIST f n` + \\ `LENGTH pi_w64_indices = 25` + by rewrite_tac[pi_w64_indices_def, LENGTH_GENLIST] + \\ simp[LIST_EQ_REWRITE] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_GENLIST, MAP_GENLIST] + \\ rw[Abbr`n`, divides_def, bool_to_bit_def] ) + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP] + \\ simp[] + \\ `LENGTH (chunks 64 (MAP bool_to_bit bs)) = 25` + by ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ simp[bool_to_bit_def, divides_def] + \\ strip_tac \\ gs[] ) + \\ conj_asm1_tac >- ( + mp_tac pi_w64_indices_bound + \\ simp[EVERY_MEM, MEM_EL, PULL_EXISTS] + \\ disch_then drule \\ simp[] ) + \\ AP_TERM_TAC + \\ DEP_REWRITE_TAC[EL_chunks] + \\ simp[NULL_GENLIST, MAP_GENLIST] + \\ conj_asm1_tac >- ( rw[Abbr`n`, NULL_LENGTH] \\ strip_tac \\ gs[] ) + \\ simp[LIST_EQ_REWRITE, LENGTH_TAKE_EQ] + \\ reverse IF_CASES_TAC + >- ( pop_assum mp_tac \\ simp[Abbr`n`] ) + \\ simp[] \\ qx_gen_tac`i` \\ strip_tac + \\ DEP_REWRITE_TAC[EL_TAKE, EL_DROP, EL_MAP, EL_GENLIST] + \\ simp[] \\ AP_TERM_TAC + \\ simp[Abbr`f`, restrict_def, DIV_LT_X] + \\ ` i DIV 64 = 0` by simp[DIV_EQ_0] + \\ rw[string_to_state_array_def, restrict_def, b2w_def] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ rewrite_tac[pi_w64_indices_def] + \\ DEP_REWRITE_TAC[EL_GENLIST] + \\ simp[] +QED + +Definition chi_w64_def: + chi_w64 (s: word64 list) = + MAPi (λi w. let y = 5 * (i DIV 5) in + w ?? (~(EL (y + ((i + 1) MOD 5)) s) && + (EL (y + ((i + 2) MOD 5)) s))) s +End + +Theorem chi_w64_thm: + state_bools_w64 bs ws ∧ + string_to_state_array bs = s ⇒ + state_bools_w64 (state_array_to_string (chi s)) (chi_w64 ws) +Proof + strip_tac + \\ gs[state_bools_w64_def] + \\ conj_asm1_tac >- rw[string_to_state_array_def, b2w_def] + \\ simp[chi_def, chi_w64_def] + \\ rewrite_tac[state_array_to_string_compute] + \\ qmatch_goalsub_abbrev_tac`5 * 5 * sw` + \\ `sw = s.w` by rw[Abbr`sw`] + \\ qunabbrev_tac`sw` \\ pop_assum SUBST_ALL_TAC + \\ qmatch_goalsub_abbrev_tac`GENLIST f n` + \\ simp[MAPi_GENLIST] + \\ `LENGTH (chunks 64 (MAP bool_to_bit bs)) = 25` + by ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ simp[bool_to_bit_def, divides_def] + \\ strip_tac \\ gs[] ) + \\ qmatch_assum_abbrev_tac`m = 25` + \\ simp[LIST_EQ_REWRITE] + \\ `bs <> []` by (strip_tac \\ gs[]) + \\ `n <> 0` by gs[Abbr`n`] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[MAP_GENLIST, NULL_GENLIST] + \\ simp[bool_to_bit_def, divides_def, Abbr`n`] ) + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP] + \\ simp[] + \\ `x DIV 5 < 5` by simp[DIV_LT_X] + \\ `5 * (x DIV 5) <= 20` by simp[] + \\ simp[] + \\ conj_asm1_tac + >- ( + `25 = 20 + 5` by simp[] + \\ pop_assum SUBST1_TAC + \\ conj_tac + \\ irule LESS_EQ_LESS_TRANS + \\ qmatch_goalsub_abbrev_tac`_ + y <= _` + \\ qexists_tac`20 + y` + \\ (reverse conj_tac >- simp[]) + \\ rewrite_tac[LT_ADD_LCANCEL] + \\ simp[Abbr`y`] ) + \\ DEP_REWRITE_TAC[word_from_bin_list_not, EL_chunks] + \\ simp[NULL_GENLIST, MAP_GENLIST] + \\ simp[NULL_LENGTH] + \\ conj_tac + >- ( + irule EVERY_TAKE + \\ irule EVERY_DROP + \\ simp[EVERY_MAP, bool_to_bit_def] + \\ rw[EVERY_MEM] \\ rw[] ) + \\ DEP_REWRITE_TAC[word_from_bin_list_and] + \\ simp[] + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ simp[] + \\ AP_TERM_TAC + \\ simp[LIST_EQ_REWRITE, LENGTH_TAKE_EQ] + \\ conj_tac >- rw[Abbr`n`] + \\ qx_gen_tac`i` \\ strip_tac + \\ simp[EL_MAP, EL_TAKE, EL_ZIP, EL_MAP2] + \\ DEP_REWRITE_TAC[EL_DROP] + \\ simp[] + \\ conj_tac >- rw[Abbr`n`] + \\ simp[EL_MAP] + \\ DEP_REWRITE_TAC[EL_MAP, EL_GENLIST] + \\ conj_tac >- rw[Abbr`n`] + \\ simp[Abbr`f`, restrict_def] + \\ qpat_x_assum`LENGTH _ = 25`kall_tac + \\ rw[string_to_state_array_def, restrict_def, b2w_def, DIV_LT_X] + \\ `i DIV 64 = 0` by simp[DIV_EQ_0] + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`c <> (~a ∧ b)` + \\ `c = EL (i + 64 * x) bs` + by ( + simp[Abbr`c`] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ qspec_then`5`mp_tac DIVISION + \\ impl_tac >- rw[] + \\ disch_then(qspec_then`x`mp_tac) + \\ simp[] ) + \\ simp[Abbr`c`] + \\ rw[bool_to_bit_def] +QED + +Definition iota_w64_RCz_def: + iota_w64_RCz : word64 list = MAP n2w [ + 1; 32898; + 9223372036854808714; 9223372039002292224; + 32907; 2147483649; + 9223372039002292353; 9223372036854808585; + 138; 136; + 2147516425; 2147483658; + 2147516555; 9223372036854775947; + 9223372036854808713; 9223372036854808579; + 9223372036854808578; 9223372036854775936; + 32778; 9223372039002259466; + 9223372039002292353; 9223372036854808704; + 2147483649; 9223372039002292232 + ] +End + +Definition iota_w64_def: + iota_w64 (s: word64 list) c = + case s of [] => [] | h :: t => (c ?? h) :: t +End + +Theorem iota_w64_RCz_rc_thm: + z < 64 /\ i < 24 ==> + EL z (PAD_RIGHT 0 64 (w2l 2 (EL i iota_w64_RCz))) = + if z = 2 ** LOG 2 (SUC z) - 1 then + bool_to_bit $ rc (7 * i + LOG 2 (SUC z)) + else 0 +Proof + strip_tac + \\ qpat_x_assum`i < _`mp_tac + \\ simp_tac std_ss [NUMERAL_LESS_THM] + \\ strip_tac \\ rpt BasicProvers.VAR_EQ_TAC \\ EVAL_TAC + \\ pop_assum mp_tac + \\ CONV_TAC(LAND_CONV(SIMP_CONV std_ss [NUMERAL_LESS_THM])) + \\ strip_tac \\ rpt BasicProvers.VAR_EQ_TAC + \\ simp_tac (srw_ss()) [bool_to_bit_def] + \\ EVAL_TAC +QED + +Theorem iota_w64_thm: + state_bools_w64 bs ws ∧ + string_to_state_array bs = s ∧ + i < 24 + ⇒ + state_bools_w64 (state_array_to_string (iota s i)) + (iota_w64 ws (EL i iota_w64_RCz)) +Proof + strip_tac + \\ gs[state_bools_w64_def] + \\ conj_asm1_tac >- rw[string_to_state_array_def, b2w_def] + \\ simp[iota_def, iota_w64_def] + \\ rewrite_tac[state_array_to_string_compute] + \\ qmatch_goalsub_abbrev_tac`5 * 5 * sw` + \\ `sw = s.w` by rw[Abbr`sw`] + \\ qunabbrev_tac`sw` \\ pop_assum SUBST_ALL_TAC + \\ qmatch_goalsub_abbrev_tac`GENLIST f n` + \\ CASE_TAC >- gs[] + \\ simp[Once chunks_def] + \\ IF_CASES_TAC >- gs[Abbr`n`] + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`t = MAP _ lg` + \\ `LENGTH lg = 24` + by ( + simp[Abbr`lg`] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ rw[bool_to_bit_def, divides_def, Abbr`n`] ) + \\ reverse conj_tac + >- ( + BasicProvers.VAR_EQ_TAC + \\ qpat_x_assum`_ = h::t`mp_tac + \\ simp[Once chunks_def] + \\ rw[] + \\ simp[LIST_EQ_REWRITE] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ rw[bool_to_bit_def, divides_def, Abbr`n`] ) + \\ simp[EL_MAP] \\ rw[Abbr`lg`] + \\ AP_TERM_TAC + \\ DEP_REWRITE_TAC[EL_chunks] + \\ simp[NULL_EQ] + \\ simp[LIST_EQ_REWRITE, LENGTH_TAKE_EQ] + \\ reverse IF_CASES_TAC >- ( + `F` suffices_by rw[] + \\ pop_assum mp_tac + \\ rw[Abbr`n`] ) + \\ simp[EL_TAKE, EL_DROP, EL_MAP] + \\ qx_gen_tac`j` \\ strip_tac + \\ simp[Abbr`f`, restrict_def, DIV_LT_X] + \\ simp[string_to_state_array_def, restrict_def, b2w_def] + \\ simp[DIV_LT_X] + \\ `j DIV 64 = 0` by simp[DIV_EQ_0] + \\ IF_CASES_TAC + >- ( + pop_assum mp_tac + \\ simp[ADD_DIV_RWT] + \\ simp[DIV_EQ_0] + \\ Cases_on`x=0` \\ simp[] + \\ Cases_on`x=1` \\ simp[] + \\ Cases_on`x=2` \\ simp[] + \\ Cases_on`x=3` \\ simp[] ) + \\ AP_TERM_TAC + \\ AP_THM_TAC + \\ AP_TERM_TAC + \\ simp[] + \\ simp[ADD_DIV_RWT] + \\ qmatch_goalsub_abbrev_tac`5 * (xx DIV 5)` + \\ qspec_then`5`mp_tac DIVISION + \\ impl_tac >- rw[] + \\ disch_then(qspec_then`xx`mp_tac) + \\ simp[] + \\ disch_then (SUBST1_TAC o SYM) + \\ simp[Abbr`xx`] ) + \\ simp[MAP_GENLIST, TAKE_GENLIST] + \\ `MIN 64 n = 64` by simp[Abbr`n`] + \\ pop_assum SUBST1_TAC + \\ qmatch_goalsub_abbrev_tac`GENLIST _ m` + \\ `h = word_from_bin_list (PAD_RIGHT 0 64 $ word_to_bin_list h)` + by simp[word_from_bin_list_def, word_to_bin_list_def, l2w_w2l] + \\ pop_assum SUBST1_TAC + \\ qmatch_goalsub_abbrev_tac`_ ?? g` + \\ `g = word_from_bin_list (PAD_RIGHT 0 64 $ word_to_bin_list g)` + by simp[word_from_bin_list_def, word_to_bin_list_def, l2w_w2l] + \\ pop_assum SUBST1_TAC + \\ DEP_REWRITE_TAC[word_from_bin_list_xor] + \\ simp[] + \\ AP_TERM_TAC + \\ BasicProvers.VAR_EQ_TAC + \\ qpat_x_assum`_ = h::t`mp_tac + \\ simp[Once chunks_def] + \\ IF_CASES_TAC >- fs[Abbr`m`] + \\ simp[] \\ strip_tac + \\ BasicProvers.VAR_EQ_TAC + \\ BasicProvers.VAR_EQ_TAC + \\ simp[LIST_EQ_REWRITE] + \\ conj_tac + >- simp[Abbr`m`] + \\ qx_gen_tac`z` + \\ qmatch_goalsub_abbrev_tac`_ ==> rsh` + \\ simp[Abbr`m`] \\ strip_tac + \\ qunabbrev_tac`rsh` + \\ DEP_REWRITE_TAC[EL_GENLIST, EL_MAP, EL_ZIP] + \\ conj_tac >- simp[] + \\ simp[] + \\ simp[word_to_bin_list_def] + \\ simp[word_from_bin_list_def, w2l_l2w] + \\ qmatch_goalsub_abbrev_tac`l2n 2 ll MOD xx` + \\ qspecl_then[`ll`,`2`]mp_tac l2n_lt + \\ simp[Abbr`ll`] + \\ strip_tac + \\ `z DIV 64 = 0` by simp[DIV_EQ_0] + \\ rw[Abbr`f`, restrict_def] + \\ rw[string_to_state_array_def, restrict_def, b2w_def] + \\ simp[iota_some_elim] + \\ DEP_REWRITE_TAC[n2l_l2n] + \\ conj_tac + >- ( + simp[] + \\ irule EVERY_TAKE + \\ simp[EVERY_MAP, EVERY_MEM] + \\ rw[bool_to_bit_def] ) + \\ simp[Abbr`g`] + \\ qmatch_goalsub_abbrev_tac`l2n 2 w` + \\ `LOG 2 (SUC z) <= w2l 64` + by ( + rewrite_tac[definition"w2l_def"] + \\ reverse IF_CASES_TAC >- rw[] + \\ rewrite_tac[LOG2_def] + \\ irule LOG2_LE_MONO + \\ simp[] ) + \\ simp[LOG2_def] + \\ simp[l2n_eq_0] + \\ qmatch_goalsub_abbrev_tac`COND b0` + \\ `¬b0 ⇒ LOG 2 (l2n 2 w) = PRE (LENGTH (dropWhile ((=) 0) (REVERSE w)))` + by ( + simp[Abbr`b0`] + \\ strip_tac + \\ DEP_REWRITE_TAC[LOG_l2n_dropWhile] + \\ pop_assum mp_tac + \\ simp[EXISTS_MEM, PULL_EXISTS, EVERY_MEM] + \\ qx_gen_tac`e` \\ strip_tac + \\ qexists_tac`e` \\ simp[] + \\ reverse conj_asm2_tac + >- ( + rw[Abbr`w`, GSYM MAP_TAKE, MEM_MAP, bool_to_bit_def] + \\ rw[] ) + \\ first_x_assum drule + \\ Cases_on`e = 0` \\ gs[] ) + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`PRE m` + \\ `¬b0 ⇒ 0 < m` + by ( + rw[Abbr`b0`, Abbr`m`] + \\ CCONTR_TAC \\ fs[dropWhile_eq_nil] + \\ rw[] + \\ ntac 2 (pop_assum mp_tac) + \\ simp[EVERY_MEM, EXISTS_MEM, PULL_EXISTS] + \\ qx_gen_tac`e` \\ strip_tac + \\ qexists_tac`e` \\ simp[] + \\ strip_tac \\ fs[] ) + \\ simp[iffLR SUC_PRE] + \\ qmatch_goalsub_abbrev_tac`_ + EL z ww` + \\ `ww = w` + by ( + simp[Abbr`ww`, Abbr`w`] + \\ simp[LIST_EQ_REWRITE, length_pad_right] + \\ conj_asm1_tac + >- ( + IF_CASES_TAC \\ simp[] + \\ IF_CASES_TAC \\ fs[] + \\ fs[Abbr`b0`] + \\ qunabbrev_tac`m` + \\ qmatch_goalsub_abbrev_tac`dropWhile P ls` + \\ qspecl_then[`P`,`ls`]mp_tac LENGTH_dropWhile_LESS_EQ + \\ simp[Abbr`ls`] + \\ simp[LESS_OR_EQ] + \\ strip_tac \\ gs[] ) + \\ simp[EL_TAKE, EL_MAP] + \\ qx_gen_tac`x` \\ strip_tac + \\ simp[PAD_RIGHT, EL_GENLIST, EL_APPEND_EQN] + \\ Cases_on`b0` \\ gs[] + >- ( + rw[bool_to_bit_def] + \\ qhdtm_x_assum`EVERY`mp_tac + \\ simp[EXISTS_MEM, MEM_EL, PULL_EXISTS] + \\ qexists_tac`x` + \\ DEP_REWRITE_TAC[EL_TAKE, EL_MAP] + \\ simp[bool_to_bit_def] ) + \\ DEP_REWRITE_TAC[TAKE_TAKE, LENGTH_TAKE] + \\ simp[] + \\ conj_asm1_tac + >- ( + reverse conj_asm1_tac >- simp[Abbr`m`] + \\ simp[Abbr`m`] + \\ qmatch_goalsub_abbrev_tac`dropWhile P ls` + \\ qspecl_then[`P`,`ls`]mp_tac LENGTH_dropWhile_LESS_EQ + \\ simp[Abbr`ls`] ) + \\ simp[EL_TAKE, EL_MAP] + \\ rw[] + \\ fs[NOT_LESS] + \\ qunabbrev_tac`m` + \\ qpat_x_assum`m <= x`(mp_then Any mp_tac EL_LENGTH_dropWhile_REVERSE) + \\ simp[EL_TAKE, EL_MAP]) + \\ fs[Abbr`ww`] + \\ simp[Abbr`w`, EL_TAKE, EL_MAP] + \\ simp[bool_to_bit_neq_add] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ AP_TERM_TAC + \\ simp[iota_w64_RCz_rc_thm] + \\ rw[bool_to_bit_def] + \\ fs[] +QED + +Definition Rnd_w64_def: + Rnd_w64 s = + iota_w64 $ + chi_w64 $ + pi_w64 $ + rho_w64 $ + theta_w64 s +End + +Theorem Rnd_w64_thm: + state_bools_w64 bs ws ∧ + string_to_state_array bs = s ∧ + i < 24 ⇒ + state_bools_w64 (state_array_to_string (Rnd s i)) + (Rnd_w64 ws (EL i iota_w64_RCz)) +Proof + rewrite_tac[Rnd_w64_def, Rnd_def] + \\ strip_tac + \\ irule iota_w64_thm + \\ rw[] + \\ qmatch_goalsub_abbrev_tac`string_to_state_array _ = s` + \\ qexists_tac`state_array_to_string s` + \\ DEP_REWRITE_TAC[state_array_to_string_to_state_array] + \\ simp[Abbr`s`] + \\ irule chi_w64_thm + \\ qmatch_goalsub_abbrev_tac`string_to_state_array _ = s` + \\ qexists_tac`state_array_to_string s` + \\ DEP_REWRITE_TAC[state_array_to_string_to_state_array] + \\ simp[Abbr`s`] + \\ irule pi_w64_thm + \\ qmatch_goalsub_abbrev_tac`string_to_state_array _ = s` + \\ qexists_tac`state_array_to_string s` + \\ DEP_REWRITE_TAC[state_array_to_string_to_state_array] + \\ simp[Abbr`s`] + \\ irule rho_w64_thm + \\ qmatch_goalsub_abbrev_tac`string_to_state_array _ = s` + \\ qexists_tac`state_array_to_string s` + \\ DEP_REWRITE_TAC[state_array_to_string_to_state_array] + \\ simp[Abbr`s`] + \\ irule theta_w64_thm + \\ metis_tac[] +QED + +Definition Keccak_p_24_w64_def: + Keccak_p_24_w64 s = + FOLDL Rnd_w64 s iota_w64_RCz +End + +Theorem Keccak_p_24_w64_lemma[local]: + state_bools_w64 bs ws /\ i <= 24 ==> + state_bools_w64 + (state_array_to_string $ + FST $ FUNPOW (λ(a,i). (Rnd a i, SUC i)) i + (string_to_state_array bs, 0)) + (FOLDL Rnd_w64 ws (TAKE i iota_w64_RCz)) +Proof + Induct_on`i` \\ rw[] + >- ( + DEP_REWRITE_TAC[GEN_ALL string_to_state_array_to_string] + \\ gs[state_bools_w64_def, iota_w64_RCz_def] ) + \\ gs[FUNPOW_SUC, UNCURRY] + \\ `LENGTH iota_w64_RCz = 24` by simp[iota_w64_RCz_def] + \\ qmatch_goalsub_abbrev_tac`FOLDL _ ws ls` + \\ qmatch_asmsub_abbrev_tac`FOLDL _ ws tl` + \\ `ls = SNOC (EL i iota_w64_RCz) tl` + by( + simp[Abbr`ls`, Abbr`tl`] + \\ simp[LIST_EQ_REWRITE, EL_TAKE, EL_APPEND] + \\ rw[] + \\ `i = x` by gs[] + \\ rw[] ) + \\ pop_assum SUBST1_TAC + \\ rewrite_tac[FOLDL_SNOC] + \\ qmatch_goalsub_abbrev_tac`Rnd s j` + \\ `j = i` + by ( + simp[Abbr`j`] + \\ qmatch_goalsub_abbrev_tac`FUNPOW f` + \\ `∀n x y. SND (FUNPOW f n (x,y)) = n + y` suffices_by simp[] + \\ Induct \\ rw[FUNPOW_SUC] + \\ rw[Abbr`f`, UNCURRY] ) + \\ rw[Abbr`j`] + \\ irule Rnd_w64_thm + \\ rw[] + \\ goal_assum(first_assum o mp_then Any mp_tac) + \\ DEP_REWRITE_TAC[state_array_to_string_to_state_array] + \\ rw[Abbr`s`] + \\ qmatch_goalsub_abbrev_tac`FUNPOW f` + \\ `∀n x y. wf_state_array x ==> wf_state_array $ FST (FUNPOW f n (x,y))` + suffices_by ( disch_then irule \\ fs[state_bools_w64_def] ) + \\ Induct \\ rw[FUNPOW_SUC] + \\ rw[Abbr`f`, UNCURRY] +QED + +Theorem Keccak_p_24_w64_thm: + state_bools_w64 bs ws ⇒ + state_bools_w64 (Keccak_p 24 bs) (Keccak_p_24_w64 ws) +Proof + rw[Keccak_p_def, Keccak_p_24_w64_def] + \\ `LENGTH bs = 1600` by fs[state_bools_w64_def] + \\ rw[string_to_state_array_w, b2w_def, definition"w2l_def"] + \\ drule Keccak_p_24_w64_lemma + \\ disch_then(qspec_then`24`mp_tac) + \\ `LENGTH iota_w64_RCz = 24` by simp[iota_w64_RCz_def] + \\ simp[TAKE_LENGTH_TOO_LONG] +QED + +Definition absorb_w64_def: + absorb_w64 Pis = + FOLDL (λSi Pi. Keccak_p_24_w64 (MAP2 word_xor Si Pi)) + (REPLICATE 25 0w) Pis +End + +Theorem absorb_w64_thm: + LIST_REL state_bools_w64 bs ws + ==> + state_bools_w64 + (FOLDL (λSi Pi. Keccak_p 24 (MAP2 (λx y. x <> y) Si Pi)) + (REPLICATE 1600 F) bs) + (absorb_w64 ws) +Proof + rw[absorb_w64_def] + \\ `REPLICATE 25 (0w:word64) = + MAP word_from_bin_list + (chunks 64 (MAP bool_to_bit (REPLICATE 1600 F)))` + by ( + simp[LIST_EQ_REWRITE] + \\ conj_asm1_tac + >- ( + DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ EVAL_TAC \\ rw[bool_to_bit_def] ) + \\ rw[EL_REPLICATE, EL_MAP] + \\ DEP_REWRITE_TAC[EL_chunks] + \\ simp[NULL_LENGTH] + \\ rw[word_from_bin_list_def, l2w_def] + \\ qmatch_goalsub_abbrev_tac`A MOD B = 0` + \\ `A = 0` suffices_by rw[] + \\ rw[Abbr`A`, l2n_eq_0] + \\ irule EVERY_TAKE + \\ rw[REPLICATE_GENLIST, bool_to_bit_def] ) + \\ pop_assum SUBST1_TAC + \\ `LENGTH (REPLICATE 1600 F) = 1600` by simp[] + \\ pop_assum mp_tac + \\ qspec_tac(`REPLICATE 1600 F`,`s0`) + \\ pop_assum mp_tac + \\ qid_spec_tac`ws` + \\ Induct_on`bs` \\ rw[] + >- rw[state_bools_w64_def] + \\ first_x_assum drule + \\ qmatch_goalsub_abbrev_tac`FOLDL _ sh bs` + \\ simp[] + \\ disch_then(qspec_then`sh`mp_tac) + \\ impl_keep_tac + >- ( + simp[Abbr`sh`, LENGTH_Keccak_p] + \\ fs[state_bools_w64_def] ) + \\ qmatch_goalsub_abbrev_tac`state_bools_w64 s1 (FOLDL f h1 ys)` + \\ strip_tac + \\ qmatch_goalsub_abbrev_tac`FOLDL f h2 ys` + \\ `h1 = h2` + by ( + simp[Abbr`h1`,Abbr`h2`, Abbr`f`, Abbr`sh`] + \\ qmatch_goalsub_abbrev_tac`Keccak_p_24_w64 ws` + \\ qmatch_goalsub_abbrev_tac`Keccak_p 24 hs` + \\ `state_bools_w64 hs ws` + by ( + simp[state_bools_w64_def, Abbr`hs`] + \\ gs[Abbr`ws`, state_bools_w64_def] + \\ simp[LIST_EQ_REWRITE] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ conj_tac >- (rpt strip_tac \\ gs[MAP2_MAP, ZIP_EQ_NIL]) + \\ simp[divides_def, bool_to_bit_def] + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP2, EL_MAP] + \\ simp[] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ conj_tac >- (rpt strip_tac \\ gs[MAP2_MAP, ZIP_EQ_NIL]) + \\ DEP_REWRITE_TAC[EL_chunks] + \\ simp[NULL_LENGTH] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_LENGTH] + \\ conj_tac >- (rpt strip_tac \\ gs[MAP2_MAP, ZIP_EQ_NIL]) + \\ conj_tac >- (rpt strip_tac \\ gs[MAP2_MAP, ZIP_EQ_NIL]) + \\ simp[word_from_bin_list_xor] + \\ AP_TERM_TAC + \\ simp[LIST_EQ_REWRITE] + \\ rpt strip_tac + \\ DEP_REWRITE_TAC[EL_MAP, EL_ZIP, EL_TAKE, EL_DROP, EL_MAP2] + \\ simp[GSYM bool_to_bit_neq_add] + \\ rw[bool_to_bit_def] \\ gs[] ) + \\ drule Keccak_p_24_w64_thm + \\ rw[state_bools_w64_def] ) + \\ rw[] +QED + +Definition eight_zeros_w64_def: + eight_zeros_w64 = [0w; 0w; 0w; 0w; 0w; 0w; 0w; 0w] : word64 list +End + +Definition Keccak_256_w64_def: + Keccak_256_w64 bs = + FLAT $ + MAP (flip word_to_bytes F) $ + TAKE 4 $ + absorb_w64 $ + pad10s1_136_w64 eight_zeros_w64 bs [] +End + +Definition Keccak_256_bytes_def: + Keccak_256_bytes (bs:word8 list) : word8 list = + MAP bools_to_word $ chunks 8 $ + Keccak_256 $ + MAP ((=) 1) $ FLAT $ + MAP (PAD_RIGHT 0 8 o word_to_bin_list) bs +End + +(* +Theorem Keccak_256_w64_thm: + Keccak_256_w64 = Keccak_256_bytes +Proof + simp[Keccak_256_w64_def, Keccak_256_bytes_def, FUN_EQ_THM] + \\ qx_gen_tac`bytes` + \\ rw[GSYM bytes_to_bools_def] + \\ rw[Keccak_256_def, Keccak_def, sponge_def] + \\ mp_tac pad10s1_136_w64_sponge_init + \\ rw[] + \\ `eight_zeros_w64 = REPLICATE 8 0w` + by rw[eight_zeros_w64_def, REPLICATE_GENLIST] + \\ pop_assum SUBST_ALL_TAC + \\ drule absorb_w64_thm \\ strip_tac + \\ qmatch_goalsub_abbrev_tac`absorb_w64 sp` + \\ qmatch_asmsub_abbrev_tac`state_bools_w64 bs` + \\ qmatch_goalsub_abbrev_tac`([], bs1)` + \\ `bs1 = bs` by rw[Abbr`bs1`, Abbr`bs`, FOLDL_MAP] + \\ gs[Abbr`bs1`] \\ pop_assum kall_tac + \\ rw[Once WHILE] + \\ rw[Once WHILE] + \\ gs[state_bools_w64_def] + \\ simp[TAKE_TAKE] + \\ simp[GSYM MAP_TAKE, MAP_MAP_o, o_DEF] + \\ simp[chunks_MAP, GSYM MAP_TAKE, MAP_MAP_o, o_DEF] + \\ DEP_REWRITE_TAC[chunks_TAKE] + \\ conj_tac >- rw[divides_def] + \\ simp[EVAL``256 DIV 8``] + \\ qmatch_goalsub_abbrev_tac`_ = MAP bw _` + \\ `bw = word_from_bin_list o MAP bool_to_bit` + by rw[bools_to_word_def, FUN_EQ_THM, Abbr`bw`] + \\ simp[Abbr`bw`, GSYM MAP_MAP_o] + \\ qmatch_goalsub_abbrev_tac`MAP f ls` + \\ `f = flip word_to_bytes F o (word_from_bin_list: num list -> word64) o + MAP bool_to_bit` by rw[Abbr`f`, FUN_EQ_THM] + \\ simp[Abbr`f`, GSYM MAP_MAP_o, Abbr`ls`] + \\ simp[MAP_TAKE, GSYM chunks_MAP] + \\ cheat +QED +*) + +(* TODO: move/replace *) +Definition hex_to_rev_bytes_def: + hex_to_rev_bytes acc [] = acc : word8 list + ∧ hex_to_rev_bytes acc [c] = CONS (n2w (UNHEX c)) acc + ∧ hex_to_rev_bytes acc (c1::c2::rest) = + hex_to_rev_bytes (CONS (n2w (16 * UNHEX c1 + UNHEX c2)) acc) rest +End + +val () = cv_auto_trans hex_to_rev_bytes_def; + +val () = cv_trans_deep_embedding EVAL eight_zeros_w64_def; + +val () = cv_auto_trans chunks_tr_aux_def; +val () = cv_auto_trans chunks_tr_def; + +val pad_pre_def = cv_auto_trans_pre (REWRITE_RULE[GSYM chunks_tr_thm]pad10s1_136_w64_def); + +Theorem pad10s1_136_w64_pre[cv_pre]: + !zs m a. pad10s1_136_w64_pre zs m a +Proof + ho_match_mp_tac pad10s1_136_w64_ind + \\ rw[] + \\ rw[Once pad_pre_def] + \\ gs[chunks_tr_thm] +QED + +Theorem theta_d_w64_inlined: + theta_d_w64 s = let + a = EL 0 s ?? EL 5 s ?? EL 10 s ?? EL 15 s ?? EL 20 s; + b = EL 1 s ?? EL 6 s ?? EL 11 s ?? EL 16 s ?? EL 21 s; + c = EL 2 s ?? EL 7 s ?? EL 12 s ?? EL 17 s ?? EL 22 s; + d = EL 3 s ?? EL 8 s ?? EL 13 s ?? EL 18 s ?? EL 23 s; + e = EL 4 s ?? EL 9 s ?? EL 14 s ?? EL 19 s ?? EL 24 s + in + [word_rol b 1 ?? e; + word_rol c 1 ?? a; + word_rol d 1 ?? b; + word_rol e 1 ?? c; + word_rol a 1 ?? d] +Proof + rw[theta_d_w64_def, theta_c_w64_def] +QED + +Theorem theta_w64_inlined: + LENGTH s = 25 ==> + theta_w64 s = let + a = EL 0 s ?? EL 5 s ?? EL 10 s ?? EL 15 s ?? EL 20 s; + b = EL 1 s ?? EL 6 s ?? EL 11 s ?? EL 16 s ?? EL 21 s; + c = EL 2 s ?? EL 7 s ?? EL 12 s ?? EL 17 s ?? EL 22 s; + d = EL 3 s ?? EL 8 s ?? EL 13 s ?? EL 18 s ?? EL 23 s; + e = EL 4 s ?? EL 9 s ?? EL 14 s ?? EL 19 s ?? EL 24 s; + j = word_rol b 1 ?? e; + k = word_rol c 1 ?? a; + l = word_rol d 1 ?? b; + m = word_rol e 1 ?? c; + n = word_rol a 1 ?? d; + in MAP2 $?? s [j;k;l;m;n;j;k;l;m;n;j;k;l;m;n;j;k;l;m;n;j;k;l;m;n] +Proof + CONV_TAC(LAND_CONV(SIMP_CONV std_ss [LENGTH_EQ_NUM_compute])) + \\ strip_tac + \\ rewrite_tac[theta_w64_def, theta_d_w64_inlined] + \\ BasicProvers.LET_ELIM_TAC + \\ gvs[Abbr`t`] +QED + +val theta_w64_pre_def = cv_auto_trans_pre (UNDISCH theta_w64_inlined); + +Theorem theta_w64_pre: + LENGTH s = 25 ==> theta_w64_pre s +Proof + rw[theta_w64_pre_def] + \\ strip_tac \\ fs[] +QED + +val rho_w64_pre_def = cv_auto_trans_pre (UNDISCH rho_w64_MAP2); + +Theorem rho_w64_pre: + LENGTH s = 25 ==> rho_w64_pre s +Proof + rw[rho_w64_pre_def] +QED + +Theorem pi_w64_inlined = SIMP_RULE std_ss [pi_w64_indices_eq, MAP] pi_w64_def; + +val pi_w64_pre_def = cv_auto_trans_pre pi_w64_inlined; + +Theorem pi_w64_pre: + LENGTH s = 25 ==> pi_w64_pre s +Proof + rw[pi_w64_pre_def] + \\ strip_tac \\ fs[] +QED + +Theorem chi_w64_inlined: + LENGTH s = 25 ==> + chi_w64 s = let + h00 = EL 0 s; h01 = EL 1 s; h02 = EL 2 s; h03 = EL 3 s; h04 = EL 4 s; + h05 = EL 5 s; h06 = EL 6 s; h07 = EL 7 s; h08 = EL 8 s; h09 = EL 9 s; + h10 = EL 10 s; h11 = EL 11 s; h12 = EL 12 s; h13 = EL 13 s; h14 = EL 14 s; + h15 = EL 15 s; h16 = EL 16 s; h17 = EL 17 s; h18 = EL 18 s; h19 = EL 19 s; + h20 = EL 20 s; h21 = EL 21 s; h22 = EL 22 s; h23 = EL 23 s; h24 = EL 24 s; + in + [h00 ⊕ h02 && ¬h01; h01 ⊕ h03 && ¬h02; h02 ⊕ h04 && ¬h03; + h03 ⊕ h00 && ¬h04; h04 ⊕ h01 && ¬h00; h05 ⊕ h07 && ¬h06; + h06 ⊕ h08 && ¬h07; h07 ⊕ h09 && ¬h08; h08 ⊕ h05 && ¬h09; + h09 ⊕ h06 && ¬h05; h10 ⊕ h12 && ¬h11; h11 ⊕ h13 && ¬h12; + h12 ⊕ h14 && ¬h13; h13 ⊕ h10 && ¬h14; h14 ⊕ h11 && ¬h10; + h15 ⊕ h17 && ¬h16; h16 ⊕ h18 && ¬h17; h17 ⊕ h19 && ¬h18; + h18 ⊕ h15 && ¬h19; h19 ⊕ h16 && ¬h15; h20 ⊕ h22 && ¬h21; + h21 ⊕ h23 && ¬h22; h22 ⊕ h24 && ¬h23; h23 ⊕ h20 && ¬h24; + h24 ⊕ h21 && ¬h20] +Proof + CONV_TAC(LAND_CONV(SIMP_CONV std_ss [LENGTH_EQ_NUM_compute])) + \\ strip_tac + \\ rewrite_tac[chi_w64_def] + \\ rw[] +QED + +val chi_w64_pre_def = cv_auto_trans_pre (UNDISCH chi_w64_inlined); + +Theorem chi_w64_pre: + LENGTH s = 25 ==> chi_w64_pre s +Proof + rw[chi_w64_pre_def] + \\ strip_tac \\ fs[] +QED + +val () = cv_auto_trans iota_w64_def; + +val Rnd_w64_pre_def = cv_auto_trans_pre Rnd_w64_def; + +Theorem Rnd_w64_pre: + LENGTH s = 25 ==> Rnd_w64_pre s w +Proof + simp[Rnd_w64_pre_def, theta_w64_pre] \\ strip_tac + \\ `LENGTH (theta_w64 s) = 25` by simp[theta_w64_inlined] + \\ simp[rho_w64_pre] + \\ `LENGTH (rho_w64 (theta_w64 s)) = 25` by ( + simp[rho_w64_MAP2] + \\ CASE_TAC \\ fs[] + \\ rw[MIN_DEF] ) + \\ simp[pi_w64_pre] + \\ irule chi_w64_pre + \\ simp[pi_w64_inlined] +QED + +Theorem Keccak_p_24_w64_inlined = + Keccak_p_24_w64_def |> SIMP_RULE std_ss [iota_w64_RCz_def, MAP, FOLDL]; + +val Keccak_p_24_w64_pre_def = cv_auto_trans_pre Keccak_p_24_w64_inlined; + +Definition absorb_w64_rec_def: + absorb_w64_rec s [] = s ∧ + absorb_w64_rec s (Pi::Pis) = + absorb_w64_rec (Keccak_p_24_w64 (MAP2 $?? s Pi)) Pis +End + +Theorem absorb_w64_rec_thm: + absorb_w64 = absorb_w64_rec (REPLICATE 25 0w) +Proof + simp[FUN_EQ_THM, absorb_w64_def, absorb_w64_rec_def] + \\ qspec_tac(`REPLICATE 25 (0w:word64)`,`s`) + \\ CONV_TAC SWAP_FORALL_CONV + \\ Induct + \\ rw[absorb_w64_rec_def] +QED + +val absorb_w64_rec_pre_def = cv_auto_trans_pre absorb_w64_rec_def; + +Theorem LENGTH_Rnd_w64: + LENGTH s = 25 ==> + LENGTH (Rnd_w64 s w) = 25 +Proof + rw[Rnd_w64_def] + \\ DEP_REWRITE_TAC[theta_w64_inlined] + \\ conj_tac >- rw[] + \\ BasicProvers.LET_ELIM_TAC + \\ DEP_REWRITE_TAC[rho_w64_MAP2] + \\ conj_tac >- simp[] + \\ CASE_TAC >- gs[LENGTH_EQ_NUM_compute] + \\ rewrite_tac[pi_w64_inlined] + \\ DEP_REWRITE_TAC[chi_w64_inlined] + \\ qmatch_goalsub_abbrev_tac`EL _ zz` + \\ conj_tac >- simp[] + \\ BasicProvers.LET_ELIM_TAC + \\ rewrite_tac[iota_w64_def] + \\ simp[] +QED + +Theorem absorb_w64_rec_pre: + ∀v s. LENGTH s = 25 /\ EVERY (((=) 25) o LENGTH) v ==> absorb_w64_rec_pre s v +Proof + Induct + \\ simp[Once absorb_w64_rec_pre_def] + \\ rpt gen_tac \\ strip_tac + \\ qmatch_goalsub_abbrev_tac`absorb_w64_rec_pre x` + \\ `LENGTH x = 25` + by ( + qunabbrev_tac`x` + \\ simp[Keccak_p_24_w64_inlined] + \\ DEP_REWRITE_TAC[LENGTH_Rnd_w64] + \\ simp[] ) + \\ rewrite_tac[Keccak_p_24_w64_pre_def] + \\ DEP_REWRITE_TAC[Rnd_w64_pre] + \\ DEP_REWRITE_TAC[LENGTH_Rnd_w64] + \\ simp[] +QED + +val Keccak_256_w64_pre_def = cv_auto_trans_pre $ + (Keccak_256_w64_def |> SIMP_RULE std_ss [C_DEF, absorb_w64_rec_thm]); + +Theorem Keccak_256_w64_pre[cv_pre]: + Keccak_256_w64_pre bytes +Proof + rw[Keccak_256_w64_pre_def] + \\ irule absorb_w64_rec_pre + \\ conj_tac >- rw[] + \\ mp_tac pad10s1_136_w64_sponge_init + \\ rw[eight_zeros_w64_def] + \\ rw[EVERY_MEM, MEM_EL] + \\ gs[LIST_REL_EL_EQN] + \\ qmatch_goalsub_abbrev_tac`pad10s1_136_w64 r8` + \\ `r8 = REPLICATE 8 0w` by simp[Abbr`r8`, REPLICATE_GENLIST] + \\ gs[] + \\ first_x_assum drule + \\ rw[state_bools_w64_def] + \\ DEP_REWRITE_TAC[LENGTH_chunks] + \\ gs[NULL_LENGTH, divides_def, bool_to_bit_def] + \\ strip_tac \\ fs[] +QED + +Theorem Keccak_256_w64_NIL: + Keccak_256_w64 [] = + REVERSE $ hex_to_rev_bytes [] + "C5D2460186F7233C927E7DB2DCC703C0E500B653CA82273B7BFAD8045D85A470" +Proof + CONV_TAC cv_eval +QED val _ = cv_trans index_to_triple_def; val _ = cv_trans triple_to_index_def; @@ -2012,12 +4117,6 @@ Proof rewrite_tac [GSYM sptreeTheory.LENGTH_toSortedAList] \\ rw [] QED -Triviality GENLIST_EQ_NIL: - GENLIST f n = [] ⇔ n = 0 -Proof - Cases_on ‘n’ \\ gvs [GENLIST] -QED - Triviality size_while1_neq_0: ∀a tt ww x y w z s. size s ≠ 0 ⇒ diff --git a/examples/Crypto/RSA/binomialScript.sml b/examples/Crypto/RSA/binomialScript.sml index 6c997dcdfc..2b8c679b32 100644 --- a/examples/Crypto/RSA/binomialScript.sml +++ b/examples/Crypto/RSA/binomialScript.sml @@ -10,9 +10,6 @@ open bossLib arithmeticTheory powerTheory summationTheory ; quietdec := false; *) -infix THEN THENC THENL; -infix 8 by; - val ARW = RW_TAC arith_ss; val _ = new_theory "binomial"; diff --git a/examples/Crypto/RSA/congruentScript.sml b/examples/Crypto/RSA/congruentScript.sml index e9504dbce0..7dd798ba08 100644 --- a/examples/Crypto/RSA/congruentScript.sml +++ b/examples/Crypto/RSA/congruentScript.sml @@ -11,9 +11,6 @@ open HolKernel Parse boolLib bossLib quietdec := false; *) -infix THEN THENC THENL; -infix 8 by; - val ARW = RW_TAC arith_ss; val _ = new_theory "congruent"; diff --git a/examples/Crypto/RSA/fermatScript.sml b/examples/Crypto/RSA/fermatScript.sml index 2df704d2b5..a72e5610f3 100644 --- a/examples/Crypto/RSA/fermatScript.sml +++ b/examples/Crypto/RSA/fermatScript.sml @@ -13,8 +13,6 @@ open bossLib numLib arithmeticTheory prim_recTheory quietdec := false; *) -infix THEN THENC THENL; -infix 8 by; val ARW = RW_TAC arith_ss; val _ = new_theory "fermat"; diff --git a/examples/Crypto/RSA/powerScript.sml b/examples/Crypto/RSA/powerScript.sml index 8fb1869d0d..9888e9d4a6 100644 --- a/examples/Crypto/RSA/powerScript.sml +++ b/examples/Crypto/RSA/powerScript.sml @@ -10,9 +10,6 @@ open HolKernel Parse boolLib quietdec := false; *) -infix THEN THENC THENL; -infix 8 by; - val _ = new_theory "power"; val ARW = RW_TAC arith_ss; diff --git a/examples/Crypto/RSA/rsaScript.sml b/examples/Crypto/RSA/rsaScript.sml index 1278eaaac4..6b24d07669 100644 --- a/examples/Crypto/RSA/rsaScript.sml +++ b/examples/Crypto/RSA/rsaScript.sml @@ -13,10 +13,6 @@ open HolKernel Parse boolLib bossLib numLib quietdec := false; *) -infix THEN THENC THENL; -infix 8 by; - - val ARW = RW_TAC arith_ss; val _ = new_theory "rsa"; diff --git a/examples/Crypto/RSA/summationScript.sml b/examples/Crypto/RSA/summationScript.sml index e3c5041688..7c6a6f5296 100644 --- a/examples/Crypto/RSA/summationScript.sml +++ b/examples/Crypto/RSA/summationScript.sml @@ -6,8 +6,6 @@ app load ["bossLib", "arithmeticTheory", "numLib"]; open bossLib arithmeticTheory numLib; -infix THEN THENC THENL; -infix 8 by; val ARW = RW_TAC arith_ss; val _ = new_theory "summation"; diff --git a/examples/Crypto/pedersenCommitment/Holmakefile b/examples/Crypto/pedersenCommitment/Holmakefile index ff60aa1f91..59137b0210 100755 --- a/examples/Crypto/pedersenCommitment/Holmakefile +++ b/examples/Crypto/pedersenCommitment/Holmakefile @@ -1,4 +1,3 @@ INCLUDES = $(HOLDIR)/examples/algebra/field \ $(HOLDIR)/src/algebra/construction \ - $(HOLDIR)/src/algebra/construction \ - $(HOLDIR)/src/num/theories \ No newline at end of file + $(HOLDIR)/src/algebra/construction diff --git a/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml b/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml index 6d45189eee..e6fd0a3855 100755 --- a/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml +++ b/examples/Crypto/pedersenCommitment/pedersenCommitmentScript.sml @@ -62,20 +62,22 @@ val GF_mult_rsub = store_thm( ‘Field (GF q)’ by rw[GF_field] >> metis_tac[field_mult_rsub]); -val field_mult_not_zero = store_thm( - "field_mult_not_zero", - “!r:'a field. Field r ==> - !d m i. d IN R ∧ m IN R ∧ i IN R ∧ d = i * m ∧ i ≠ #0 ∧ m ≠ #0 ==> (d ≠ #0)”, +Theorem field_mult_not_zero: + !r:'a field. + Field r ==> + !d m i. d IN R ∧ m IN R ∧ i IN R ∧ d = i * m ∧ i ≠ #0 ∧ m ≠ #0 ==> (d ≠ #0) +Proof rpt strip_tac >> rw[] >> - ‘i = #0 ∨ m = #0’ by metis_tac[field_mult_eq_zero]); + ‘i = #0 ∨ m = #0’ by metis_tac[field_mult_eq_zero] +QED val GF_mult_not_zero = store_thm( "GF_mult_not_zero", “∀q. prime q ⇒ ∀d m i. - m IN (GF q).carrier ∧ - d IN (GF q).carrier ∧ + m IN (GF q).carrier ∧ + d IN (GF q).carrier ∧ i IN (GF q).carrier ∧ d = i _*_ m ∧ i ≠ _0_ ∧ @@ -88,7 +90,7 @@ val GF_add_sub_identity = store_thm( "GF_add_sub_identity", “∀q. prime q ⇒ ∀x y z t. - x IN (GF q).carrier ∧ + x IN (GF q).carrier ∧ y IN (GF q).carrier ∧ z IN (GF q).carrier ∧ t IN (GF q).carrier ∧ @@ -102,7 +104,7 @@ val GF_diff_mult_solve = store_thm( "GF_diff_mult_solve", “∀q. prime q ⇒ ∀d1 d2 m1 m2 i. - d1 IN (GF q).carrier ∧ + d1 IN (GF q).carrier ∧ d2 IN (GF q).carrier ∧ m1 IN (GF q).carrier ∧ m2 IN (GF q).carrier ∧ @@ -128,12 +130,12 @@ End (* Theorems involving commit and verify *) val GF_pedersen_binding = store_thm( "GF_pedersen_binding", - “∀q. prime q ⇒ + “∀q. prime q ⇒ ∀m1 m2 d1 d2. m1 IN (GF q).carrier ∧ m2 IN (GF q).carrier ∧ d1 IN (GF q).carrier ∧ - d2 IN (GF q).carrier ∧ + d2 IN (GF q).carrier ∧ (m1 ≠ m2) ⇒ ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ ∀k. k IN G ∧ @@ -170,12 +172,12 @@ Theorem GF_pedersen_binding_prime = GF_pedersen_binding |> SIMP_RULE (srw_ss()) val GF_pedersen_hiding = store_thm( "GF_pedersen_hiding", - “∀q. prime q ⇒ - ∀m. m IN (GF q).carrier ⇒ - ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ + “∀q. prime q ⇒ + ∀m. m IN (GF q).carrier ⇒ + ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ ∀c k. c IN G ∧ k IN G ⇒ ∃d. d IN (GF q).carrier ∧ - c = commit g (cyclic_gen g) d k m”, + c = commit g (cyclic_gen g) d k m”, simp[commit_def] >> rw[] >> ‘∃i. i < CARD G ∧ c = monoid_exp g (cyclic_gen g) i’ by rw[cyclic_element_by_gen] >> @@ -196,12 +198,12 @@ val GF_pedersen_hiding = store_thm( val ZN_pedersen_hiding = store_thm( "ZN_pedersen_hiding", - “∀q. prime q ⇒ - ∀m. m IN (ZN q).carrier ⇒ - ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ + “∀q. prime q ⇒ + ∀m. m IN (ZN q).carrier ⇒ + ∀g: 'a group. cyclic g ∧ FINITE G ∧ (ord (cyclic_gen g) = q) ⇒ ∀c k. c IN G ∧ k IN G ⇒ ∃d. d IN (ZN q).carrier ∧ - c = commit g (cyclic_gen g) d k m”, + c = commit g (cyclic_gen g) d k m”, simp[commit_def] >> rw[] >> ‘∃i. i < CARD G ∧ c = monoid_exp g (cyclic_gen g) i’ by rw[cyclic_element_by_gen] >> diff --git a/examples/Crypto/sigmaProtocol/Holmakefile b/examples/Crypto/sigmaProtocol/Holmakefile new file mode 100755 index 0000000000..e1344c18de --- /dev/null +++ b/examples/Crypto/sigmaProtocol/Holmakefile @@ -0,0 +1,6 @@ +INCLUDES = $(HOLDIR)/examples/Crypto/pedersenCommitment \ + $(HOLDIR)/examples/computability/register \ + $(HOLDIR)/examples/set-theory/vbg \ + $(HOLDIR)/src/algebra/construction + + diff --git a/examples/Crypto/sigmaProtocol/sigmaProtocolScript.sml b/examples/Crypto/sigmaProtocol/sigmaProtocolScript.sml new file mode 100644 index 0000000000..bb1772840f --- /dev/null +++ b/examples/Crypto/sigmaProtocol/sigmaProtocolScript.sml @@ -0,0 +1,1339 @@ + +(* ------------------------------------------------------------------------- *) +(* Theory: Sigma Protocol *) +(* *) +(* This theory defines and proves properties of Sigma protocols and their *) +(* combiners. *) +(* The key components are: *) +(* *) +(* - Abstract Sigma Protocol definition with Statements, Witnesses, *) +(* Relation, Commitments, Challenges, and other core functionality *) +(* *) +(* - Protocol Combiners: *) +(* - OR combiner for proving knowledge of one of two statements *) +(* - AND combiner for proving knowledge of both statements *) +(* - Equality combiner for proving same witness works for two statements *) +(* *) +(* - Properties proven for protocols and combiners: *) +(* - Well-formedness *) +(* - Completeness *) +(* - Special Soundness *) +(* - Simulator Correctness *) +(* - Honest Verifier Zero-Knowledge *) +(* *) +(* ------------------------------------------------------------------------- *) + +(* Import External Theories *) +open HolKernel boolLib bossLib Parse; +open jcLib; +open vbgsetTheory; +open rmToPairTheory; +open ringTheory; +open groupTheory; +open arithmeticTheory; + + +(* Declare new theory at start *) +val _ = new_theory "sigmaProtocol"; + +(* Define a symbol for Field sum operation *) +Overload "⊕" = “(GF q).sum.op”; +val _ = set_fixity "⊕" (Infixl 500); + +(* Define a symbol for Field subtraction operation *) +Overload "⊖" = “ring_sub (GF q)”; +val _ = set_fixity "⊖" (Infixl 500); + +(* Define a symbol for Field multiplication operation *) +Overload "⊗" = “(GF q).prod.op”; +val _ = set_fixity "⊗" (Infixl 600); + +(* Define a symbol for Field division operation *) +Overload "⊘" = “field_div (GF q)”; +val _ = set_fixity "⊘" (Infixl 600); + + +(* Definition of the cross poduct of the groups *) +Definition Gcross_def: + Gcross g1 g2 = <| carrier:= g1.carrier × g2.carrier; + id:= (g1.id, g2.id); + op:= λ (a,b) (c,d). (g1.op a c, g2.op b d); + |> +End + + +(* Cross product of Abelian groups is an Abelian group *) +val Gcross_group_thm = store_thm( + "Gcross_group_thm", + “AbelianGroup g1 ∧ AbelianGroup g2 ⇒ AbelianGroup (Gcross g1 g2)”, + rw[] >> + simp[AbelianGroup_def, group_def_alt, Group_def, Gcross_def, pairTheory.FORALL_PROD] >> + rpt (strip_tac >> + gvs[AbelianGroup_def, group_op_element]) >> + rpt (metis_tac[AbelianGroup_def, group_assoc]) >> + (‘(∀x1. x1 ∈ g1.carrier ⇒ ∃y1. y1 ∈ g1.carrier ∧ g1.op y1 x1 = g1.id) ∧ + (∀x2. x2 ∈ g2.carrier ⇒ ∃y2. y2 ∈ g2.carrier ∧ g2.op y2 x2 = g2.id)’ by metis_tac[AbelianGroup_def, group_def_alt] >> + ‘(∃y1. y1 ∈ g1.carrier ∧ g1.op y1 p_1 = g1.id) ∧ + (∃y2. y2 ∈ g2.carrier ∧ g2.op y2 p_2 = g2.id)’ by metis_tac[AbelianGroup_def, group_def_alt] >> + qabbrev_tac‘y = (y1, y2)’ >> + ‘(λ(a,b) (c,d). (g1.op a c,g2.op b d)) (y1, y2) (p_1,p_2) = ((g1.op y1 p_1) , (g2.op y2 p_2))’ by rw[] >> + ‘((g1.op y1 p_1) , (g2.op y2 p_2)) = (g1.id,g2.id)’ by metis_tac[pairTheory.PAIR] >> + qexists_tac ‘y’ >> + ‘FST y ∈ g1.carrier ∧ SND y ∈ g2.carrier’ by rw[pairTheory.PAIR, Abbr‘y’] >> + metis_tac[])); + + +(* Definition of an Abstract Sigma Protocol - + a three-move interactive system where the Prover convinces the Verifier of knowledge + of a secret (witness) for some statement, while not revealing (zero-knowledge) anything else about the secret *) +Datatype: SigmaProtocol = <| + + (* Set of all possible statements that can be proven *) + Statements: 's set; + + (* Set of all possible witnesses - secret values that prove the statements *) + Witnesses: 'w set; + + (* Binary relation defining which witnesses prove which statements + Returns true if witness w proves statement s *) + Relation: 's -> 'w -> bool; + + (* Set of random values used by the prover to maintain zero-knowledge property *) + RandomCoins: 'r set; + + (* Set of all possible first messages (commitments to the witness) from prover to verifier. + Maintain binding property *) + Commitments: 'c set; + + (* Algebraic group of all possible challenge values sent by verifier to prover. + Typically this is a finite cyclic group *) + Challenges: 'e group; + + (* Binary relation defining when two challenges are considered distinct + Used in the special soundness property *) + Disjoint: 'e -> 'e -> bool; + + (* Set of all possible responses from prover to verifier's challenge *) + Responses: 't set; + + (* First prover algorithm that generates the initial commitment + Takes statement s, witness w, and random coins r + The witness is needed to ensure the commitment is properly bound to the proof *) + Prover_0: 's -> 'w -> 'r -> 'c; + + (* Second prover algorithm that generates the response to the challenge + Takes statement s, witness w, random coins r, commitment c, and challenge e *) + Prover_1: 's -> 'w -> 'r -> 'c -> 'e -> 't; + + (* Verifier algorithm that checks if a proof transcript is valid + Takes a tuple of (statement, commitment, challenge, response) + Returns true if the proof is accepted *) + HonestVerifier: ('s # 'c # 'e # 't) -> bool; + + (* Knowledge extractor that can compute a valid witness given two accepting transcripts + with the same commitment but different challenges + This ensures proof of knowledge property *) + Extractor: 't -> 't -> 'e -> 'e -> 'w; + + (* Simulator that can generate valid-looking transcripts without knowing the witness + This ensures zero-knowledge property *) + Simulator: 's -> 't -> 'e -> ('s # 'c # 'e # 't); + + (* Maps witness and randomness to simulated response + Used to show equivalence between real and simulated proofs *) + SimulatorMap: 's -> 'w -> 'e -> 'r -> 't; + + (* Inverse of SimulatorMap - recovers randomness from simulated response + Used to show simulator is perfect *) + SimulatorMapInverse: 's -> 'w -> 'e -> 't -> 'r; + |> +End + + +(* Helper definition of the Difference Operation between two challenges in the challenge group + Used to ensure challenge consistency in the OR composition of the protocol *) +Definition SP_csub_def: + SP_csub sp a b = sp.Challenges.op a (sp.Challenges.inv b) +End + + +(* OR composition of a sigma protocol that allows proving knowledge of a witness + for at least one of two statements without revealing which one *) +Definition SP_or_def: + SP_or sp = <| + (* Statement pairs - prover will prove knowledge for at least one *) + Statements:= sp.Statements × sp.Statements; + + (* Witness type remains the same - only need witness for one statement *) + Witnesses:= sp.Witnesses; + + (* New relation - satisfied if witness proves either statement *) + Relation:= (λ (s1, s2) w. sp.Relation s1 w ∨ sp.Relation s2 w); + + (* Random coins now include: + - Original random coins for real proof + - Response for simulated proof + - Challenge for simulated proof *) + RandomCoins:= ((sp.RandomCoins × sp.Responses) × sp.Challenges.carrier); + + (* Need commitments for both statements *) + Commitments:= (sp.Commitments × sp.Commitments); + + (* Challenge space remains the same *) + Challenges:= sp.Challenges; + + (* Two challenges are disjoint if they're not equal *) + Disjoint:= (λ e1 e2. e1 ≠ e2); + + (* Response now includes: + - Response for first statement + - Challenge used in simulation + - Response for second statement *) + Responses:= ((sp.Responses × sp.Challenges.carrier) × sp.Responses); + + (* First prover algorithm that generates initial commitments + If we have witness for first statement: + - Generate real commitment for first statement + - Simulate proof for second statement + If we have witness for second statement: + - Simulate proof for first statement + - Generate real commitment for second statement *) + Prover_0:= (λ (s1, s2) w ((r, t),e). + if sp.Relation s1 w then + let c1 = sp.Prover_0 s1 w r; + (s2', c2', e2', t2') = sp.Simulator s2 t e; + in (c1, c2') + else + let (s1', c1', e1', t1') = sp.Simulator s1 t e; + c2 = sp.Prover_0 s2 w r; + in (c1',c2)); + + (* Second prover algorithm that generates responses + e3 is computed as e1 - e2 to ensure challenge consistency + If we have witness for first statement: + - Use simulator for second statement + - Generate real response for first statement with challenge e3 + If we have witness for second statement: + - Use simulator for first statement + - Generate real response for second statement with challenge e3 *) + Prover_1:= (λ (s1, s2) w ((r,t1),e2) (c1, c2) e1. + let e3 = SP_csub sp e1 e2 + in + if sp.Relation s1 w then + let (s2', c2', e2', t2') = sp.Simulator s2 t1 e2; + t2 = sp.Prover_1 s1 w r c1 e3; + in ((t2, e3), t2') + else + let (s1', c1', e1', t1') = sp.Simulator s1 t1 e2; + t2 = sp.Prover_1 s2 w r c2 e3; + in ((t1', e2), t2)); + + (* Verifier checks both proofs are valid and challenges sum correctly *) + HonestVerifier:= (λ ((s1, s2), (c1, c2), e1, ((t1,e2),t2)). + let e3 = SP_csub sp e1 e2 + in sp.HonestVerifier (s1, c1, e2, t1) ∧ + sp.HonestVerifier (s2, c2, e3, t2)); + + (* Knowledge extractor that can recover witness from two accepting transcripts + If first challenges differ, extract from first statement responses + Otherwise extract from second statement responses using computed e2, e4 *) + Extractor:= (λ ((t1,e1),t2) ((t3,e3),t4) e5 e6. + if e1 ≠ e3 then + sp.Extractor t1 t3 e1 e3 + else + let e2 = SP_csub sp e5 e1; + e4 = SP_csub sp e6 e3; + in sp.Extractor t2 t4 e2 e4); + + (* Simulator generates accepting transcripts for both statements + Uses original simulator twice and ensures challenge consistency *) + Simulator:= (λ (s1, s2) ((t1,e1),t2) e2. + let (s1', c1', e1', t1') = sp.Simulator s1 t1 e1; + e3 = SP_csub sp e2 e1; + (s2', c2', e2', t2') = sp.Simulator s2 t2 e3; + in ((s1, s2),(c1', c2'),e2,((t1',e1), t2'))); + + (* Maps witness and randomness to simulated response + If proving first statement: simulate second, map first + If proving second statement: simulate first, map second *) + SimulatorMap:= (λ (s1, s2) w e1 ((r,t),e2). + let e3 = SP_csub sp e1 e2; + in + if sp.Relation s1 w then + let t1 = sp.SimulatorMap s1 w e3 r; + in ((t1, e3), t) + else + let t2 = sp.SimulatorMap s2 w e3 r; + in ((t, e2), t2)); + + (* Inverse of SimulatorMap - recovers randomness + Determines which statement the witness proves and + inverts the appropriate simulation *) + SimulatorMapInverse:= (λ (s1, s2) w e1 ((t1,e2),t2). + let e3 = SP_csub sp e1 e2; + in + if sp.Relation s1 w then + let r = sp.SimulatorMapInverse s1 w e2 t1; + in ((r,t2),e3) + else + let r = sp.SimulatorMapInverse s2 w e3 t2; + in ((r,t1),e2)); + |> +End + + +(* AND composition of two sigma protocols S1 and S2 that allows proving knowledge + of witnesses for both statements simultaneously. This creates a compound sigma protocol + where both component proofs must be valid *) +Definition SP_and_def: + SP_and S1 S2 = <| + (* Statements are pairs - must prove both *) + Statements:= S1.Statements × S2.Statements; + + (* Need witnesses for both statements *) + Witnesses:= S1.Witnesses × S2.Witnesses; + + (* Relation holds only if both component relations hold *) + Relation:= (λ (s1, s2) (w1, w2). S1.Relation s1 w1 ∧ S2.Relation s2 w2); + + (* Need random coins for both proofs *) + RandomCoins:= (S1.RandomCoins × S2.RandomCoins); + + (* Need commitments for both proofs *) + Commitments:= (S1.Commitments × S2.Commitments); + + (* Challenge space is the direct product of both challenge groups *) + Challenges:= Gcross S1.Challenges S2.Challenges; + + (* Two challenge pairs are disjoint if they're disjoint in both components *) + Disjoint:= λ (a,b) (c,d). S1.Disjoint a c ∧ S2.Disjoint b d; + + (* Need responses for both proofs *) + Responses:= (S1.Responses × S2.Responses); + + (* First prover algorithm runs both component Prover_0s independently *) + Prover_0:= (λ (s1, s2) (w1, w2) (r1, r2). + (S1.Prover_0 s1 w1 r1, S2.Prover_0 s2 w2 r2)); + + (* Second prover algorithm runs both component Prover_1s independently + with their respective challenges *) + Prover_1:= (λ (s1, s2) (w1, w2) (r1, r2) (c1, c2) (e1, e2). + (S1.Prover_1 s1 w1 r1 c1 e1, S2.Prover_1 s2 w2 r2 c2 e2)); + + (* Verifier accepts only if both component verifiers accept *) + HonestVerifier:= (λ ((s1, s2), (c1, c2), (e1, e2), (t1, t2)). + S1.HonestVerifier (s1, c1, e1, t1) ∧ + S2.HonestVerifier (s2, c2, e2, t2)); + + (* Extractor runs both component extractors independently *) + Extractor:= (λ (t1a, t1b) (t2a, t2b) (e1a, e1b) (e2a, e2b). + (S1.Extractor t1a t2a e1a e2a, S2.Extractor t1b t2b e1b e2b)); + + (* Simulator runs both component simulators independently and combines results *) + Simulator:= (λ (s1, s2) (t1, t2) (e1, e2). + let (s1', c1', e1', t1') = S1.Simulator s1 t1 e1; + (s2', c2', e2', t2') = S2.Simulator s2 t2 e2 + in + ((s1', s2'),(c1', c2'),(e1', e2'),(t1', t2'))); + + (* SimulatorMap runs both component maps independently *) + SimulatorMap:= (λ (s1, s2) (w1, w2) (e1, e2) (r1, r2). + (S1.SimulatorMap s1 w1 e1 r1, S2.SimulatorMap s2 w2 e2 r2)); + + (* SimulatorMapInverse runs both component inverses independently *) + SimulatorMapInverse:= (λ (s1, s2) (w1, w2) (e1, e2) (t1, t2). + (S1.SimulatorMapInverse s1 w1 e1 t1, + S2.SimulatorMapInverse s2 w2 e2 t2)); + |> +End + + +(* Equality composition of a sigma protocol S1 that proves a witness satisfies + the same relation for two different statements. This creates a compound protocol + that proves knowledge of a single witness that works for both statements. *) +Definition SP_eq_def: + SP_eq S1 = <| + (* Statement pairs - must prove both with same witness *) + Statements:= S1.Statements × S1.Statements; + + (* Single witness must work for both statements *) + Witnesses:= S1.Witnesses; + + (* Relation holds if witness proves both statements *) + Relation:= (λ (s1, s2) w. S1.Relation s1 w ∧ S1.Relation s2 w); + + (* Use same random coins for both proofs to ensure consistency *) + RandomCoins:= S1.RandomCoins; + + (* Need commitments for both statements *) + Commitments:= (S1.Commitments × S1.Commitments); + + (* Use challenge space from original protocol *) + Challenges:= S1.Challenges; + + (* Use same disjointness relation *) + Disjoint:= S1.Disjoint; + + (* Single response proves both statements *) + Responses:= S1.Responses; + + (* First prover algorithm generates commitments for both statements + using same randomness to ensure consistency *) + Prover_0:= (λ (s1, s2) w r. (S1.Prover_0 s1 w r, S1.Prover_0 s2 w r)); + + (* Second prover algorithm generates single response that works for both + statements since we used same randomness *) + Prover_1:= (λ (s1, s2) w r (c1, c2) e. + (S1.Prover_1 s1 w r c1 e)); + + (* Verifier checks same response validates both proofs *) + HonestVerifier:= (λ ((s1, s2), (c1, c2), e, t). + S1.HonestVerifier (s1, c1, e, t) ∧ + S1.HonestVerifier (s2, c2, e, t)); + + (* Can use original extractor since responses are the same *) + Extractor:= S1.Extractor; + + (* Simulator generates consistent transcripts for both statements *) + Simulator:= (λ (s1, s2) t e. + let (s1', c1', e1', t1') = S1.Simulator s1 t e; + (s2', c2', e2', t2') = S1.Simulator s2 t e + in + ((s1', s2'),(c1', c2'),e1', t1')); + + (* Can use original simulator map since responses are the same *) + SimulatorMap:= (λ (s1, s2) w e r. + (S1.SimulatorMap s1 w e r)); + + (* Can use original simulator map inverse *) + SimulatorMapInverse:= (λ (s1, s2) w e t. + (S1.SimulatorMapInverse s1 w e t)); + |> +End + + +(* Disjointness relation in the equality composition + is exactly the same as in the original protocol. *) +val sp_eq_dis_thm = store_thm( + "sp_eq_dis_thm", + “(SP_eq sp).Disjoint = sp.Disjoint”, + rw[] >> + simp[SP_eq_def]); + + +(* Well-formedness conditions for a sigma protocol, + particularly focused on requirements needed for the equality combiner *) +Definition WellFormed_SP_def: + WellFormed_SP sp ⇔ + ( + (* Challenge space must form an abelian group *) + (AbelianGroup sp.Challenges) ∧ + + (* Disjoint challenges must be different values + This ensures soundness of the protocol *) + (∀e1 e2. e1 ∈ sp.Challenges.carrier ∧ e2 ∈ sp.Challenges.carrier ⇒ + (sp.Disjoint e1 e2 ⇒ (e1 ≠ e2))) ∧ + + (* First prover message must produce valid commitments + given well-typed inputs *) + (∀s w r. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ⇒ + sp.Prover_0 s w r ∈ sp.Commitments) ∧ + + (* Second prover message must produce valid responses + given well-typed inputs *) + (∀s w r c e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ∧ + c ∈ sp.Commitments ∧ e ∈ sp.Challenges.carrier ⇒ + sp.Prover_1 s w r c e ∈ sp.Responses) ∧ + + (* Extractor must produce valid witnesses + given valid responses and challenges *) + (∀t1 t2 e1 e2. t1 ∈ sp.Responses ∧ t2 ∈ sp.Responses ∧ + e1 ∈ sp.Challenges.carrier ∧ e2 ∈ sp.Challenges.carrier ⇒ + sp.Extractor t1 t2 e1 e2 ∈ sp.Witnesses) ∧ + + (* Simulator must: + - Preserve the input statement + - Preserve the input challenge + - Preserve the input response + - Produce valid commitments *) + (∀s t e s' c' e' t'. + s ∈ sp.Statements ∧ t ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier ∧ + sp.Simulator s t e = (s', c', e', t') ⇒ + s' = s ∧ + e' = e ∧ + t' = t ∧ + c' ∈ sp.Commitments) ∧ + + (* SimulatorMap must produce valid responses + given well-typed inputs *) + (∀s w r e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMap s w e r ∈ sp.Responses) ∧ + + (* SimulatorMapInverse must produce valid random coins + given well-typed inputs *) + (∀s w t e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + t ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMapInverse s w e t ∈ sp.RandomCoins) + ) +End + + +(* Well-formedness conditions for a sigma protocol + specifically for use with the OR combiner. Similar to WellFormed_SP but with + an additional requirement about the disjointness relation. *) +Definition Or_WellFormed_SP_def: + Or_WellFormed_SP sp ⇔ + ( + (* Challenge space must form an abelian group *) + (AbelianGroup sp.Challenges) ∧ + + (* Disjoint challenges must be different values *) + (∀e1 e2. e1 ∈ sp.Challenges.carrier ∧ e2 ∈ sp.Challenges.carrier ⇒ + (sp.Disjoint e1 e2 ⇒ (e1 ≠ e2))) ∧ + + (* First prover message must produce valid commitments + given well-typed inputs *) + (∀s w r. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ⇒ + sp.Prover_0 s w r ∈ sp.Commitments) ∧ + + (* Second prover message must produce valid responses + given well-typed inputs *) + (∀s w r c e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ∧ + c ∈ sp.Commitments ∧ e ∈ sp.Challenges.carrier ⇒ + sp.Prover_1 s w r c e ∈ sp.Responses) ∧ + + (* Extractor must produce valid witnesses + given valid responses and challenges *) + (∀t1 t2 e1 e2. t1 ∈ sp.Responses ∧ t2 ∈ sp.Responses ∧ + e1 ∈ sp.Challenges.carrier ∧ e2 ∈ sp.Challenges.carrier ⇒ + sp.Extractor t1 t2 e1 e2 ∈ sp.Witnesses) ∧ + + (* Simulator must: + - Preserve the input statement + - Preserve the input challenge + - Preserve the input response + - Produce valid commitments *) + (∀s t e s' c' e' t'. + s ∈ sp.Statements ∧ t ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier ∧ + sp.Simulator s t e = (s', c', e', t') ⇒ + s' = s ∧ + e' = e ∧ + t' = t ∧ + c' ∈ sp.Commitments) ∧ + + (* SimulatorMap must produce valid responses + given well-typed inputs *) + (∀s w r e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMap s w e r ∈ sp.Responses) ∧ + + (* SimulatorMapInverse must produce valid random coins + given well-typed inputs *) + (∀s w t e. s ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + t ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMapInverse s w e t ∈ sp.RandomCoins) ∧ + + (* Additional requirement for OR combiner: + Disjointness is exactly inequality - challenges are disjoint + if and only if they are different. This is needed to ensure + proper challenge manipulation in the OR proof. *) + (∀e1 e2. e1 ∈ sp.Challenges.carrier ∧ e2 ∈ sp.Challenges.carrier ⇒ + (sp.Disjoint e1 e2 ⇔ e1 ≠ e2)) + ) +End + + +(* Well-formedness conditions specifically required for equality composition of sigma protocols. + These extend the basic well-formedness with additional requirements ensuring responses + and simulation are independent of the statement being proven. *) +Definition Eq_WellFormed_SP_def: + Eq_WellFormed_SP sp ⇔ + ( + (* Protocol must satisfy basic well-formedness conditions *) + WellFormed_SP sp ∧ + + (* Prover_1 must: + - Produce valid responses + - Generate the same response regardless of which statement is being proven + This ensures the same witness can be used for both statements *) + (∀s1 s2 w r c1 c2 e. s1 ∈ sp.Statements ∧ + s2 ∈ sp.Statements ∧ + w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ + c1 ∈ sp.Commitments ∧ + c2 ∈ sp.Commitments ∧ + e ∈ sp.Challenges.carrier ⇒ + sp.Prover_1 s1 w r c1 e ∈ sp.Responses ∧ + sp.Prover_1 s1 w r c1 e = sp.Prover_1 s2 w r c2 e) ∧ + + (* SimulatorMap must: + - Produce valid responses for both statements + - Generate the same response regardless of statement + This ensures consistent simulation across both statements *) + (∀s1 s2 w r e. s1 ∈ sp.Statements ∧ + s2 ∈ sp.Statements ∧ + w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ + e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMap s1 w e r ∈ sp.Responses ∧ + sp.SimulatorMap s2 w e r ∈ sp.Responses ∧ + sp.SimulatorMap s1 w e r = sp.SimulatorMap s2 w e r ) ∧ + + (* SimulatorMapInverse must: + - Produce valid random coins for both statements + - Generate the same random coins regardless of statement + This ensures inverse simulation is also statement-independent *) + (∀s1 s2 w t e. s1 ∈ sp.Statements ∧ + s2 ∈ sp.Statements ∧ + w ∈ sp.Witnesses ∧ + t ∈ sp.Responses ∧ + e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMapInverse s1 w e t ∈ sp.RandomCoins ∧ + sp.SimulatorMapInverse s2 w e t ∈ sp.RandomCoins ∧ + sp.SimulatorMapInverse s1 w e t = sp.SimulatorMapInverse s2 w e t) + ) +End + + +(* Completeness property for sigma protocols: For any valid statement-witness pair, + an honest prover following the protocol will always convince an honest verifier. + This is a fundamental requirement for any proof system. *) +Definition Complete_SP_def: + Complete_SP sp ⇔ + (* For all valid inputs: *) + (∀s w r e. + (* Given well-typed parameters *) + s ∈ sp.Statements ∧ + w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ + e ∈ sp.Challenges.carrier ∧ + (* And given that witness proves the statement *) + sp.Relation s w ⇒ + (* Let c be the prover's first message (commitment) *) + let c = sp.Prover_0 s w r in + (* Then the verifier must accept the complete transcript where: + - s is the statement being proven + - c is the prover's commitment + - e is the verifier's challenge + - Prover_1 generates the response based on all previous values *) + sp.HonestVerifier(s, c, e, sp.Prover_1 s w r c e)) +End + + +(* Special Soundness property for sigma protocols: given two accepting transcripts + with the same commitment but different challenges, we can extract a valid witness. + This property ensures knowledge soundness - the prover must "know" the witness. *) +Definition SpecialSoundness_SP_def: + SpecialSoundness_SP sp ⇔ + ∀s c e1 e2 t1 t2. + (* Given well-typed parameters *) + s ∈ sp.Statements ∧ + c ∈ sp.Commitments ∧ + t1 ∈ sp.Responses ∧ + t2 ∈ sp.Responses ∧ + e1 ∈ sp.Challenges.carrier ∧ + e2 ∈ sp.Challenges.carrier ∧ + (* And given two accepting transcripts where: + - Same statement and commitment + - Different (disjoint) challenges + - Both verify correctly *) + sp.Disjoint e1 e2 ∧ + sp.HonestVerifier(s, c, e1, t1) ∧ + sp.HonestVerifier(s, c, e2, t2) ⇒ + (* Then the extracted witness must be valid for the statement *) + sp.Relation s (sp.Extractor t1 t2 e1 e2) +End + + +(* Simulator Correctness property: simulated transcripts must be accepted by + the honest verifier. This is a basic requirement for zero-knowledge simulation + - simulated proofs should look valid. *) +Definition SimulatorCorrectness_SP_def: + SimulatorCorrectness_SP sp ⇔ + ∀s t e. + (* Given well-typed parameters *) + s ∈ sp.Statements ∧ + t ∈ sp.Responses ∧ + e ∈ sp.Challenges.carrier ⇒ + (* Simulated transcript must verify *) + sp.HonestVerifier(sp.Simulator s t e) +End + + +(* Honest Verifier Zero Knowledge property: for any valid statement-witness pair, + real proofs are indistinguishable from simulated ones. This ensures the proof + reveals nothing about the witness beyond its validity. *) +Definition HonestVerifierZeroKnowledge_SP_def: + HonestVerifierZeroKnowledge_SP sp ⇔ + ∀s w r e t. + (* Given well-typed parameters and valid statement-witness pair *) + s ∈ sp.Statements ∧ + w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ + e ∈ sp.Challenges.carrier ∧ + t ∈ sp.Responses ∧ + sp.Relation s w ⇒ + let c = sp.Prover_0 s w r; + spm = sp.SimulatorMap s w e; + spmi = sp.SimulatorMapInverse s w e; + in + (* Three key properties: + 1. Simulation matches real proofs exactly + 2. SimulatorMapInverse undoes SimulatorMap + 3. SimulatorMap undoes SimulatorMapInverse + Together these show perfect simulation *) + sp.Simulator s (spm r) e = (s, c, e, sp.Prover_1 s w r c e) ∧ + spmi(spm r) = r ∧ + spm(spmi t) = t +End + + +(* AND composition of well-formed sigma protocols is also well-formed. + This shows the AND combiner preserves well-formedness. *) +val WellFormed_SP_and_thm = store_thm( + "WellFormed_SP_and_thm", + “WellFormed_SP sp1 ∧ WellFormed_SP sp2 ⇒ WellFormed_SP (SP_and sp1 sp2)”, + simp[WellFormed_SP_def, SP_and_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >- + ‘AbelianGroup (Gcross sp1.Challenges sp2.Challenges)’ by metis_tac[Gcross_group_thm] >> + rpt (pairarg_tac >> gvs[]) >> + fs[Gcross_def, pairTheory.PAIR] >> + gs[WellFormed_SP_def, pairTheory.PAIR] >> + metis_tac[pairTheory.PAIR]); + + +(* AND composition of complete sigma protocols is also complete. + Shows that the AND combiner preserves the ability for honest provers to convince honest verifiers. *) +val Complete_SP_and_thm = store_thm( + "Complete_SP_and_thm", + “Complete_SP sp1 ∧ Complete_SP sp2 ⇒ Complete_SP (SP_and sp1 sp2)”, + simp[Complete_SP_def, SP_and_def, pairTheory.FORALL_PROD] >> + rpt (strip_tac >> gs[] >> fs[Gcross_def, pairTheory.PAIR])); + + +(* AND composition of special sound sigma protocols is also special sound. + Shows that the AND combiner preserves the ability to extract witnesses + from accepting transcripts with different challenges. *) +val SpecialSoundness_SP_and_thm = store_thm( + "SpecialSoundness_SP_and_thm", + “SpecialSoundness_SP sp1 ∧ SpecialSoundness_SP sp2 ⇒ SpecialSoundness_SP (SP_and sp1 sp2)”, + simp[SpecialSoundness_SP_def, SP_and_def, pairTheory.FORALL_PROD] >> + rpt (strip_tac >> fs[Gcross_def, pairTheory.PAIR] >> metis_tac[])); + + +(* AND composition of simulator-correct sigma protocols is also simulator-correct. + Shows that the AND combiner preserves the ability to generate valid-looking + simulated transcripts that pass verification. *) +val SimulatorCorrectness_SP_and_thm = store_thm( + "SimulatorCorrectness_SP_and_thm", + “SimulatorCorrectness_SP sp1 ∧ SimulatorCorrectness_SP sp2 ⇒ SimulatorCorrectness_SP (SP_and sp1 sp2)”, + simp[SimulatorCorrectness_SP_def, SP_and_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + ‘p_1'' ∈ sp1.Challenges.carrier ∧ p_2'' ∈ sp2.Challenges.carrier’ by fs[Gcross_def, pairTheory.PAIR] >> + rpt (pairarg_tac >> gvs[]) >> + metis_tac[]); + + +(* AND composition of Honest-Verifier Zero-Knowledge (HVZK) sigma protocols is also HVZK, + given the component protocols are well-formed. This shows the AND combiner + preserves zero-knowledge simulation capabilities. *) +val HonestVerifierZeroKnowledge_SP_and_thm = store_thm( + "HonestVerifierZeroKnowledge_SP_and_thm", + “HonestVerifierZeroKnowledge_SP sp1 ∧ HonestVerifierZeroKnowledge_SP sp2 ∧ + WellFormed_SP sp1 ∧ WellFormed_SP sp2 ⇒ HonestVerifierZeroKnowledge_SP (SP_and sp1 sp2)”, + simp[HonestVerifierZeroKnowledge_SP_def, SP_and_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >- + (qabbrev_tac‘s1 = p_1’ >> + qabbrev_tac‘s2 = p_2’ >> + qabbrev_tac‘w1 = p_1'’ >> + qabbrev_tac‘w2 = p_2'’ >> + qabbrev_tac‘r1 = p_1''’ >> + qabbrev_tac‘r2 = p_2''’ >> + qabbrev_tac‘e1 = p_1'³'’ >> + qabbrev_tac‘e2 = p_2'³'’ >> + qabbrev_tac‘t1 = p_1'⁴'’ >> + qabbrev_tac‘t2 = p_2'⁴'’ >> + ‘e1 ∈ sp1.Challenges.carrier ∧ e2 ∈ sp2.Challenges.carrier’ by fs[Gcross_def, pairTheory.PAIR] >> + ‘sp1.SimulatorMapInverse s1 w1 e1 (sp1.SimulatorMap s1 w1 e1 r1) = r1 ∧ + sp1.SimulatorMap s1 w1 e1 (sp1.SimulatorMapInverse s1 w1 e1 t1) = t1 ∧ + sp2.SimulatorMapInverse s2 w2 e2 (sp2.SimulatorMap s2 w2 e2 r2) = r2 ∧ + sp2.SimulatorMap s2 w2 e2 (sp2.SimulatorMapInverse s2 w2 e2 t2) = t2’ by metis_tac[] >> + ‘WellFormed_SP (SP_and sp1 sp2)’ by metis_tac[WellFormed_SP_and_thm] >> + ‘sp1.Prover_0 s1 w1 r1 ∈ sp1.Commitments ∧ sp2.Prover_0 s2 w2 r2 ∈ sp2.Commitments’ by metis_tac[WellFormed_SP_def] >> + qabbrev_tac‘c1 = sp1.Prover_0 s1 w1 r1’ >> qabbrev_tac‘c2 = sp2.Prover_0 s2 w2 r2’ >> + qabbrev_tac‘t11 = sp1.Prover_1 s1 w1 r1 c1 e1’ >> + qabbrev_tac‘t22 = sp2.Prover_1 s2 w2 r2 c2 e2’ >> + qabbrev_tac‘t222 = sp2.SimulatorMap s2 w2 e2 r2’ >> + qabbrev_tac‘t111 = sp1.SimulatorMap s1 w1 e1 r1’ >> + qabbrev_tac‘x1 = sp1.Simulator s1 t111 e1’ >> + qabbrev_tac‘x2 = sp2.Simulator s2 t222 e2’ >> + ‘(sp1.Prover_1 s1 w1 r1 c1 e1) ∈ sp1.Responses’ by metis_tac[WellFormed_SP_def] >> + ‘t11 ∈ sp1.Responses’ by metis_tac[] >> + ‘t11 ∈ sp1.Responses ∧ t111 ∈ sp1.Responses ∧ t22 ∈ sp2.Responses ∧ t222 ∈ sp2.Responses’ by metis_tac[WellFormed_SP_def] >> + ‘sp1.Simulator s1 t111 e1 = sp1.Simulator s1 (sp1.SimulatorMap s1 w1 e1 r1) e1’ by rw[] >> + ‘sp1.Simulator s1 (sp1.SimulatorMap s1 w1 e1 r1) e1 = + (s1,sp1.Prover_0 s1 w1 r1, e1, sp1.Prover_1 s1 w1 r1 (sp1.Prover_0 s1 w1 r1) e1)’ by metis_tac[] >> + ‘(s1,sp1.Prover_0 s1 w1 r1, e1, sp1.Prover_1 s1 w1 r1 (sp1.Prover_0 s1 w1 r1) e1) = (s1, c1, e1, t11)’ by metis_tac[] >> + ‘x1 = (s1, c1, e1, t11)’ by gs[] >> + + ‘sp2.Simulator s2 t222 e2 = + sp2.Simulator s2 (sp2.SimulatorMap s2 w2 e2 r2) e2’ by rw[] >> + ‘sp2.Simulator s2 (sp2.SimulatorMap s2 w2 e2 r2) e2 = + (s2,sp2.Prover_0 s2 w2 r2, e2, sp2.Prover_1 s2 w2 r2 (sp2.Prover_0 s2 w2 r2) e2)’ by metis_tac[] >> + ‘(s2,sp2.Prover_0 s2 w2 r2, e2, sp2.Prover_1 s2 w2 r2 (sp2.Prover_0 s2 w2 r2) e2) = (s2, c2, e2, t22)’ by metis_tac[] >> + ‘x2 = (s2, c2, e2, t22)’ by metis_tac[] >> + ‘(λ(s1',c1',e1',t1').(λ(s2',c2',e2',t2'). ((s1',s2'),(c1',c2'),(e1',e2'),t1',t2')) (s2, c2, e2, t22)) (s1, c1, e1, t11) = + ((s1,s2),(c1,c2),(e1,e2),(t11,t22))’ by rw[] >> + ‘((s1,s2),(c1,c2),(e1,e2),t11,t22) = (λ(s1',c1',e1',t1'). (λ(s2',c2',e2',t2'). ((s1',s2'),(c1',c2'),(e1',e2'),t1',t2')) x2) x1’ by metis_tac[] >> + metis_tac[]) >> + rpt(gs[Gcross_def, pairTheory.PAIR] >> metis_tac[])); + + +(* Lemma proving SP_csub (challenge subtraction) produces valid challenges. + This is crucial for OR composition where we need to manipulate challenges + while ensuring they remain in the challenge group's carrier set. *) +val SP_csub_Lemma = store_thm( + "SP_csub_Lemma", + “∀sp x y. + WellFormed_SP sp ∧ + x ∈ sp.Challenges.carrier ∧ + y ∈ sp.Challenges.carrier ⇒ + SP_csub sp x y ∈ sp.Challenges.carrier”, + simp[SP_csub_def, SP_or_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + qabbrev_tac‘y1 = sp.Challenges.inv y’ >> + qabbrev_tac‘Gr = sp.Challenges’ >> + ‘Group Gr’ by metis_tac[AbelianGroup_def, WellFormed_SP_def] >> + ‘y1 ∈ Gr.carrier’ by metis_tac[group_inv_element] >> + ‘Gr.op x y1 ∈ Gr.carrier’ by metis_tac[group_op_element]); + + +(* Auxiliary tactics *) +val CASES_ON_WITNESS = +(fn () => + Cases_on ‘sp1.Relation p_1 w’ >> + rpt ( + pairarg_tac >> + gvs[] + ) >> + fs[] >> + metis_tac[WellFormed_SP_def, SP_or_def, SP_csub_Lemma] >> + fs[] +) (); + + +val PAIR_TAC_METIS_TAC = +(fn () => + rpt (pairarg_tac>> gvs[]) >> + metis_tac[WellFormed_SP_def, SP_or_def, SP_csub_Lemma] +) (); + + +(* OR composition preserves well-formedness - if the input protocol sp1 + is well-formed, then its OR composition SP_or sp1 is also well-formed *) +val WellFormed_SP_or_thm = store_thm( + "WellFormed_SP_or_thm", + “WellFormed_SP sp1 ⇒ WellFormed_SP (SP_or sp1)”, + rw[]>> + simp[WellFormed_SP_def, SP_or_def, pairTheory.FORALL_PROD]>> + rpt strip_tac>- + PAIR_TAC_METIS_TAC >> + rpt(CASES_ON_WITNESS)>- + (Cases_on ‘p_2 ≠ p_2'' ’ >> + rpt(fs[] >> metis_tac[WellFormed_SP_def, SP_or_def, SP_csub_Lemma]))>> + rpt(PAIR_TAC_METIS_TAC)>> + rpt(CASES_ON_WITNESS)); + + +(* OR composition preserves simulator correctness - if the input protocol sp1 + has correct simulation and is well-formed, then its OR composition SP_or sp1 also produces + valid simulated transcripts that pass verification. *) +val SimulatorCorrectness_SP_or_thm = store_thm( + "SimulatorCorrectness_SP_or_thm", + “SimulatorCorrectness_SP sp1 ∧ WellFormed_SP sp1 ⇒ SimulatorCorrectness_SP (SP_or sp1)”, + simp[SimulatorCorrectness_SP_def, SP_or_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + qabbrev_tac‘s1 = p_1’ >> qabbrev_tac‘s2 = p_2’ >> qabbrev_tac‘t1 = p_1'’ >> + qabbrev_tac‘e1 = p_2'’ >> qabbrev_tac‘t2 = p_2''’ >> + qabbrev_tac‘e2 = (SP_csub sp1 e e1)’ >> qabbrev_tac‘c1 = sp1.Prover_0 s1 w r1’ >> + ‘e2 ∈ sp1.Challenges.carrier’ by metis_tac[SP_csub_Lemma] >> + qabbrev_tac‘t3 = sp1.Prover_1 s1 w r1 c1 e2’ >> + qabbrev_tac‘sa = FST (sp1.Simulator s1 t1 e1)’ >> + qabbrev_tac‘ca = FST (SND (sp1.Simulator s1 t1 e1))’ >> + qabbrev_tac‘ea = FST (SND (SND (sp1.Simulator s1 t1 e1)))’ >> + qabbrev_tac‘ta = SND (SND (SND (sp1.Simulator s1 t1 e1)))’ >> + ‘sp1.Simulator s1 t1 e1 = (sa, ca, ea, ta)’ by rw[pairTheory.PAIR, Abbr‘sa’, Abbr‘ca’, Abbr‘ea’, Abbr‘ta’] >> + qabbrev_tac‘sb = FST (sp1.Simulator s2 t2 e2)’ >> + qabbrev_tac‘cb = FST (SND (sp1.Simulator s2 t2 e2))’ >> + qabbrev_tac‘eb = FST (SND (SND (sp1.Simulator s2 t2 e2)))’ >> + qabbrev_tac‘tb = SND (SND (SND (sp1.Simulator s2 t2 e2)))’ >> + ‘sp1.Simulator s2 t2 e2 = (sb, cb, eb, tb)’ by rw[pairTheory.PAIR, Abbr‘sb’, Abbr‘cb’, Abbr‘eb’, Abbr‘tb’] >> + gs[] >> + metis_tac[WellFormed_SP_def]); + + +(* Lemma proving a key property of challenge subtraction (SP_csub) in the challenge group: + For any challenges x,y in a well-formed protocol's challenge space, we have + SP_csub sp x (SP_csub sp x y) = y + This is like the algebraic identity: + x - (x-y) = x - x + y = y *) +val SP_csub_Lemma2 = store_thm( + "SP_csub_Lemma2", + “∀sp x y z. + WellFormed_SP sp ∧ + x ∈ sp.Challenges.carrier ∧ + y ∈ sp.Challenges.carrier ⇒ + SP_csub sp x (SP_csub sp x y) = y”, + simp[SP_csub_def, SP_or_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + qabbrev_tac‘Gr = sp.Challenges’ >> + ‘AbelianGroup Gr’ by metis_tac[WellFormed_SP_def] >> + ‘Group Gr’ by metis_tac[AbelianGroup_def] >> + ‘Gr.inv y ∈ Gr.carrier’ by metis_tac[group_inv_element] >> + ‘Gr.op x (Gr.inv y) ∈ Gr.carrier’ by metis_tac[group_op_element] >> + qabbrev_tac‘z = Gr.inv (Gr.op x (Gr.inv y))’ >> + qabbrev_tac‘s = Gr.inv y’ >> + ‘Gr.inv s = y’ by metis_tac[group_inv_inv] >> + ‘Gr.op x s = Gr.op s x’ by metis_tac[AbelianGroup_def] >> + ‘z = Gr.inv (Gr.op s x)’ by metis_tac[] >> + ‘Gr.op Gr.id y = y’ by metis_tac[group_lid] >> + ‘Gr.op x (Gr.inv x) = Gr.id’ by metis_tac[group_rinv] >> + ‘Gr.op (Gr.op x (Gr.inv x)) y = y’ by metis_tac[] >> + ‘Gr.inv x ∈ Gr.carrier’ by metis_tac[group_inv_element] >> + ‘Gr.op x (Gr.op (Gr.inv x) y) = y’ by metis_tac[group_linv_assoc] >> + ‘Gr.op x (Gr.op (Gr.inv x) (Gr.inv s)) = y’ by metis_tac[] >> + ‘Gr.inv (Gr.op s x) = Gr.op (Gr.inv x) (Gr.inv s)’ by metis_tac[group_inv_op] >> + ‘Gr.op x (Gr.inv (Gr.op s x)) = y’ by metis_tac[] >> + ‘Gr.op x z = y’ by metis_tac[Abbr‘z’]); + + +(* Lemma proving another key property of challenge subtraction (SP_csub) in the challenge group: + If e1 ≠ e2 and e3 = e4, then SP_csub sp e1 e3 ≠ SP_csub sp e2 e4 + In other words, if we subtract the same challenge (e3 = e4) from different challenges (e1 ≠ e2), + we get different results. This is like the algebraic identity: + If a ≠ b then a - c ≠ b - c *) +val SP_csub_Lemma3 = store_thm( + "SP_csub_Lemma3", + “∀sp e1 e2 e3 e4. + WellFormed_SP sp ∧ + e1 ∈ sp.Challenges.carrier ∧ + e2 ∈ sp.Challenges.carrier ∧ + e3 ∈ sp.Challenges.carrier ∧ + e4 ∈ sp.Challenges.carrier ∧ + e1 ≠ e2 ∧ e3 = e4 + ⇒ (SP_csub sp e1 e3) ≠ (SP_csub sp e2 e4)”, + rw[] >> + simp[WellFormed_SP_def, SP_csub_def, SP_or_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + qabbrev_tac‘Gr = sp.Challenges’ >> + ‘AbelianGroup Gr’ by metis_tac[WellFormed_SP_def] >> + ‘Group Gr’ by metis_tac[AbelianGroup_def] >> + ‘e1 = Gr.op (Gr.op e2 (Gr.inv e3)) (Gr.inv (Gr.inv e3))’ by metis_tac[group_lsolve, group_op_element, group_inv_element] >> + ‘e1 = Gr.op (Gr.op e2 (Gr.inv e3)) e3’ by metis_tac[group_inv_inv] >> + ‘e1 = Gr.op e2 (Gr.op (Gr.inv e3) e3)’ by metis_tac[group_assoc, group_op_element, group_inv_element] >> + ‘e1 = Gr.op e2 Gr.id’ by metis_tac[group_linv, group_op_element, group_inv_element] >> + ‘e1 = e2’ by metis_tac[group_rid, group_op_element, group_inv_element]); + + +(* OR composition preserves completeness - if the input protocol sp1 + is complete, simulator-correct and well-formed, then its OR composition SP_or sp1 is also complete. + This shows an honest prover can convince an honest verifier in the OR composition when they + have a valid witness for either statement. *) +val Complete_SP_or_thm = store_thm( + "Complete_SP_or_thm", + “Complete_SP sp1 ∧ SimulatorCorrectness_SP sp1 ∧ WellFormed_SP sp1 ⇒ Complete_SP (SP_or sp1)”, + rw[]>> + simp[Complete_SP_def, SP_or_def, pairTheory.FORALL_PROD]>> + rpt strip_tac >- + (rpt (pairarg_tac>> + gvs[])>> + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r1 = p_1' ’>> + qabbrev_tac‘t1 = p_2' ’>> + qabbrev_tac‘e1 = p_2'' ’>> + qabbrev_tac‘e2 = (SP_csub sp1 e e1)’>> qabbrev_tac‘c1 = sp1.Prover_0 s1 w r1’>> + qabbrev_tac‘t3 = sp1.Prover_1 s1 w r1 c1 e2’>> + ‘e2 ∈ sp1.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + drule_then assume_tac (cj 1 (iffLR Complete_SP_def))>> + first_assum$ qspecl_then [‘s1’, ‘w’, ‘r1’, ‘e2’] mp_tac>> + ‘sp1.HonestVerifier (s1,c1,e2,t3)’by metis_tac[]>> + fs[]>> + ‘s2 = s2' ∧ e1 = e2'’ by metis_tac[WellFormed_SP_def]>> + ‘sp1.Simulator s2 t1 e1 = (s2,c2,e1,t2)’by rw[]>> + ‘SP_csub sp1 e e2 = e1’by metis_tac[SP_csub_Lemma2]>> + metis_tac[SimulatorCorrectness_SP_def])>- + (rpt (pairarg_tac>> + gvs[])>> + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r1 = p_1' ’>> + qabbrev_tac‘t1 = p_2' ’>> + qabbrev_tac‘e1 = p_2'' ’>> + Cases_on ‘sp1.Relation s1 w’ >| [ + ‘((sp1.Prover_1 s1 w r1 c1 (SP_csub sp1 e e1),SP_csub sp1 e e1),t2') = ((t1,e2),t2)’by metis_tac[]>> + ‘(sp1.Prover_0 s1 w r1,c2') = (c1, c2)’by metis_tac[]>> + fs[]>> + ‘c1 = sp1.Prover_0 s1 w r1’ by metis_tac[Gcross_def, pairTheory.PAIR]>> + ‘e2 = SP_csub sp1 e e1’by metis_tac[Gcross_def,pairTheory.PAIR]>> + ‘e2 ∈ sp1.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + ‘t1 = sp1.Prover_1 s1 w r1 c1 e2’by metis_tac[Gcross_def,pairTheory.PAIR]>> + ‘sp1.HonestVerifier (s1,c1,e2,t1)’by metis_tac[Complete_SP_def]>> + ‘SP_csub sp1 e e2 = e1’by metis_tac[SP_csub_Lemma2]>> + ‘s2 = s2' ∧ e1 = e2'’ by metis_tac[WellFormed_SP_def]>> (* Simulator does not change s and e *) + ‘t2 = t2' ∧ c2 = c2'’by metis_tac[Gcross_def,pairTheory.PAIR]>> + ‘ sp1.HonestVerifier (s2,c2,SP_csub sp1 e e2,t2)’by metis_tac[SimulatorCorrectness_SP_def]>> + fs[], + ‘(c1',sp1.Prover_0 s2 w r1) = (c1,c2)’by metis_tac[]>> + ‘((t1',e1),sp1.Prover_1 s2 w r1 c2 (SP_csub sp1 e e1)) = ((t1,e2),t2)’by metis_tac[]>> + fs[]>> + ‘c1' = c1 ∧ c2 = sp1.Prover_0 s2 w r1 ∧ t1' = t1 ∧ e1 = e2 ∧ sp1.Prover_1 s2 w r1 c2 (SP_csub sp1 e e1) = t2’ + by metis_tac[Gcross_def,pairTheory.PAIR]>> + ‘s1 = s1' ∧ s2 = s2'∧ e1 = e1' ∧ e1 = e2'’ by metis_tac[WellFormed_SP_def]>> + ‘sp1.HonestVerifier (s1,c1,e2,t1)’by metis_tac[SimulatorCorrectness_SP_def]>> + qabbrev_tac‘e3 = SP_csub sp1 e e1 ’>> + ‘e3 = SP_csub sp1 e e2’by rw[]>> + ‘e3 ∈ sp1.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + ‘sp1.HonestVerifier (s2,c2,e3,t2)’by metis_tac[Complete_SP_def]>> + gs[] + ])); + + +(* OR composition preserves special soundness - given a well-formed + protocol sp1 with special soundness and challenges that are disjoint iff unequal, + its OR composition SP_or sp1 also satisfies special soundness. + Special soundness means we can extract a valid witness from any two accepting transcripts + with the same commitment but different challenges. *) +val SpecialSoundness_SP_or_thm = store_thm( + "SpecialSoundness_SP_or_thm", + “∀sp. WellFormed_SP sp1 ∧ + SpecialSoundness_SP sp1 ∧ + (∀e1 e2. e1 ∈ sp1.Challenges.carrier ∧ + e2 ∈ sp1.Challenges.carrier ⇒ + (sp1.Disjoint e1 e2 ⇔ e1 ≠ e2)) + ⇒ SpecialSoundness_SP (SP_or sp1)”, + rw[] >> + simp[SpecialSoundness_SP_def, SP_or_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + qabbrev_tac‘s1 = p_1’ >> qabbrev_tac‘s2 = p_2’ >> + qabbrev_tac‘c1 = p_1'’ >> qabbrev_tac‘c2 = p_2'’ >> + qabbrev_tac‘t1 = p_1''’ >> qabbrev_tac‘t2 = p_2'³'’ >> qabbrev_tac‘t3 = p_1'³'’ >> qabbrev_tac‘t4 = p_2'⁵'’ >> + qabbrev_tac‘e3 = p_2''’ >> qabbrev_tac‘e4 = p_2''''’ >> + qabbrev_tac‘e5 = SP_csub sp1 e1 e3’ >> qabbrev_tac‘e6 = SP_csub sp1 e2 e4’ >> + ‘e5 ∈ sp1.Challenges.carrier ∧ e6 ∈ sp1.Challenges.carrier’ by metis_tac[SP_csub_Lemma] >> + Cases_on ‘e3 ≠ e4’ THENL [ + qabbrev_tac‘w = (if e3 ≠ e4 then sp1.Extractor t1 t3 e3 e4 + else sp1.Extractor t2 t4 e5 e6)’ >> + ‘w = sp1.Extractor t1 t3 e3 e4’ by metis_tac[] >> + ‘sp1.Relation s1 w’ by metis_tac[SpecialSoundness_SP_def] >> + fs[] >> + rw[], + qabbrev_tac ‘w = (if sp1.Disjoint e3 e4 then sp1.Extractor t1 t3 e3 e4 + else sp1.Extractor t2 t4 e5 e6)’ >> + ‘w = sp1.Extractor t2 t4 e5 e6’ by metis_tac[] >> + ‘e3 = e4’ by metis_tac[] >> + ‘e5 ≠ e6’ by metis_tac[SP_csub_Lemma3] >> + ‘sp1.Disjoint e5 e6’ by metis_tac[] >> + ‘sp1.Relation s2 w’ by metis_tac[SpecialSoundness_SP_def] >> + rw[] + ]); + + +(* OR composition preserves honest-verifier zero-knowledge (HVZK) +property. When the input protocol is HVZK and well-formed, its OR composition remains HVZK *) +val HonestVerifierZeroKnowledge_SP_or_thm = store_thm( + "HonestVerifierZeroKnowledge_SP_or_thm", + “HonestVerifierZeroKnowledge_SP sp ∧ WellFormed_SP sp ⇒ HonestVerifierZeroKnowledge_SP (SP_or sp)”, + rw[]>> + simp[HonestVerifierZeroKnowledge_SP_def, SP_or_def, pairTheory.FORALL_PROD]>> + rpt strip_tac >- + (rpt (pairarg_tac>> + gvs[])>> + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r = p_1' ’>> qabbrev_tac‘t1 = p_2' ’>> + qabbrev_tac‘e1 = p_2'' ’>> + qabbrev_tac‘t2 = p_1'' ’>> + qabbrev_tac‘e2 = p_2'³'’>> + qabbrev_tac‘t3 = p_2'''' ’>> + qabbrev_tac‘e3 = SP_csub sp e e1’>> + ‘e3 ∈ sp.Challenges.carrier’by metis_tac[SP_csub_Lemma]>> + qabbrev_tac‘t4 = sp.SimulatorMap s1 w e3 r’>> + ‘ t4 ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + ‘SP_csub sp e e3 = e1’by metis_tac[SP_csub_Lemma2]>> + ‘ sp.Simulator s2 t1 e1 = (s2',c2',e2',t2')’by metis_tac[]>> + ‘s1 = s1' ∧ t4 = t1' ∧ e3 = e1'’by metis_tac[WellFormed_SP_def]>> + ‘ s2 = s2'' ∧ t1 = t2'' ∧ e1 = e2''’by metis_tac[WellFormed_SP_def]>> + ‘ s2 = s2' ∧ t1 = t2' ∧ e1 = e2'’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + first_assum$ qspecl_then [‘s1’, ‘w’, ‘r’, ‘e3’, ‘t4’] mp_tac>> + fs[])>- + + (fs[]>> + qabbrev_tac‘e3 = SP_csub sp e p_2''’>> + ‘e3 ∈ sp.Challenges.carrier’by metis_tac[SP_csub_Lemma]>> + ‘p_2'' = SP_csub sp e e3’by metis_tac[SP_csub_Lemma2]>> + qabbrev_tac‘t4 = sp.SimulatorMap p_1 w e3 p_1'’>> + ‘ t4 ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘p_1'’, ‘e3’, ‘t4’] mp_tac>> + fs[])>- + (fs[]>> + qabbrev_tac‘e3 = SP_csub sp e p_2'³'’>> + ‘e3 ∈ sp.Challenges.carrier’by metis_tac[SP_csub_Lemma]>> + ‘SP_csub sp e e3 = p_2'³'’by metis_tac[SP_csub_Lemma2]>> + qabbrev_tac‘r = sp.SimulatorMapInverse p_1 w p_2'³' p_1''’>> + ‘ r ∈ sp.RandomCoins’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘r’, ‘p_2'³'’, ‘ p_1''’] mp_tac>> + fs[])>- + (rpt (pairarg_tac>> + gvs[])>> + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r = p_1' ’>> + qabbrev_tac‘t3 = p_2' ’>> + qabbrev_tac‘e3 = p_2'' ’>> + qabbrev_tac‘t4 = p_1'' ’>> + qabbrev_tac‘e4 = p_2'³'’>> + qabbrev_tac‘t5 = p_2'''' ’>> + qabbrev_tac‘e5 = SP_csub sp e e1’>> + qabbrev_tac‘e6 = SP_csub sp e e3’>> + Cases_on ‘sp.Relation s1 w’ THENL [ + ‘((sp.SimulatorMap s1 w e6 r,e6),t3) = ((t1,e1),t2) ∧ (sp.Prover_0 s1 w r,c2'') = (c1,c2)’by metis_tac[]>> + fs[]>> + ‘ t1 = sp.SimulatorMap s1 w e6 r ∧ e6 = e1 ∧ t3 = t2 ∧ sp.Prover_0 s1 w r = c1 ∧ c2'' = c2’by metis_tac[pairTheory.PAIR]>> + ‘e5 ∈ sp.Challenges.carrier ∧ e6 ∈ sp.Challenges.carrier’by metis_tac[SP_csub_Lemma, WellFormed_SP_def]>> + ‘e5 = e3’ by metis_tac[SP_csub_Lemma2]>> + ‘s1 = s1' ∧ t1 = t1' ∧ e1 = e1'’by metis_tac[WellFormed_SP_def]>> + ‘s2 = s2' ∧ t2 = t2' ∧ e5 = e2'’by metis_tac[WellFormed_SP_def]>> + ‘s2 = s2'' ∧ t3 = t2'' ∧ e3 = e2''’by metis_tac[WellFormed_SP_def]>> + ‘s1 = s1'' ∧ t3 = t1'' ∧ e3 = e1''’by metis_tac[WellFormed_SP_def]>> + fs[]>> + gvs[]>> + qabbrev_tac‘c = sp.Prover_0 p_1 w p_1'’>> + ‘ c ∈ sp.Commitments’by metis_tac[WellFormed_SP_def]>> + qabbrev_tac‘t = sp.Prover_1 p_1 w p_1' c e1’>> + ‘t ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘p_1'’, ‘e1’, ‘t’] mp_tac>> + fs[] + , + ‘(c1'',sp.Prover_0 s2 w r) = (c1,c2)’by metis_tac[]>> + fs[]>> + ‘c1'' = c1 ∧ c2 = sp.Prover_0 s2 w r’by metis_tac[WellFormed_SP_def, pairTheory.PAIR]>> + ‘((t3,e3),sp.SimulatorMap s2 w e6 r) = ((t1,e1),t2)’by metis_tac[]>> + ‘t3 = t1 ∧ e3 = e1 ∧ t2 = sp.SimulatorMap s2 w e6 r’by metis_tac[pairTheory.PAIR]>> + qabbrev_tac‘t6 = sp.Prover_1 s2 w r c2 e6’>> + ‘e6 ∈ sp.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + ‘ c2 ∈ sp.Commitments’by metis_tac[WellFormed_SP_def]>> + ‘t6 ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + ‘e5 = e6’ by metis_tac[]>> + ‘sp.Simulator s1 t1 e1 = sp.Simulator s1 t3 e3’by metis_tac[]>> + ‘(s1',c1',e1',t1') = (s1'',c1'',e1'',t1'')’by metis_tac[]>> + ‘c1' = c1’by rw[]>> + ‘s1'' = s1' ∧ s1 = s1' ∧ t1 = t1' ∧ t1' = t1'' ∧ e1 = e1' ∧ e1'' = e1'’by metis_tac[pairTheory.PAIR, WellFormed_SP_def]>> + ‘t2' ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + ‘sp.Simulator s2 t2 e6 = (s2, c2, e6, t6)’by metis_tac[]>> + ‘e5 = e6’by rw[]>> + ‘sp.Simulator s2 t2 e5 = sp.Simulator s2 t2 e6’by rw[]>> + ‘(s2',c2',e2',t2') = (s2, c2, e6, t6)’by metis_tac[pairTheory.PAIR, pairTheory.PAIR]>> + ‘c2 = c2'’by rw[pairTheory.PAIR]>> + ‘t6 = t2' ’by metis_tac[pairTheory.PAIR, WellFormed_SP_def]>> + rw[] + ])>- + (rpt (pairarg_tac>> + gvs[])>> + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r = p_1' ’>> qabbrev_tac‘t3 = p_2' ’>> + qabbrev_tac‘e1 = p_2'' ’>> + qabbrev_tac‘t4 = p_1'' ’>> + qabbrev_tac‘e3 = p_2'³'’>> + qabbrev_tac‘t5 = p_2'''' ’>> + qabbrev_tac‘e4 = SP_csub sp e e1’>> + qabbrev_tac‘e5 = SP_csub sp e e2’>> + ‘e4 ∈ sp.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + Cases_on ‘sp.Relation s1 w’ THENL [ + ‘((sp.SimulatorMap s1 w e4 r,e4),t3) = ((t1,e2),t2)’by metis_tac[]>> + ‘t1 = sp.SimulatorMap s1 w e4 r ∧ e2 = e4 ∧ t3 = t2’ by rw[pairTheory.PAIR]>> + ‘e5 = e1’by metis_tac[SP_csub_Lemma2]>> + ‘e2 ∈ sp.Challenges.carrier ∧ t1 ∈ sp.Responses’by metis_tac[WellFormed_SP_def]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + ‘sp.SimulatorMapInverse s1 w e2 t1 = r’ by metis_tac[]>> + fs[] + , + + ‘((t3,e1),sp.SimulatorMap s2 w e4 r) = ((t1,e2),t2)’by metis_tac[]>> + fs[]>> + metis_tac[pairTheory.PAIR, pairTheory.PAIR, HonestVerifierZeroKnowledge_SP_def] + ])>- + (rpt (pairarg_tac>> + gvs[])>> + Cases_on ‘sp.Relation p_1 w’ THENL [ + ‘((sp.SimulatorMapInverse p_1 w p_2'³' p_1'',p_2'⁴'), SP_csub sp e p_2'³') = ((r,t),e2)’by metis_tac[]>> + ‘r = sp.SimulatorMapInverse p_1 w p_2'³' p_1'' ∧ t = p_2'⁴' ∧ e2 = SP_csub sp e p_2'³'’by rw[pairTheory.PAIR]>> + ‘SP_csub sp e e2 ∈ sp.Challenges.carrier’by metis_tac[SP_csub_Lemma]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + ‘SP_csub sp e e2 = SP_csub sp e (SP_csub sp e p_2'³')’by rw[]>> + ‘SP_csub sp e (SP_csub sp e p_2'³') = p_2'³'’by metis_tac[SP_csub_Lemma2]>> + ‘sp.SimulatorMap p_1 w p_2'³' (sp.SimulatorMapInverse p_1 w p_2'³' p_1'') = p_1''’ by metis_tac[HonestVerifierZeroKnowledge_SP_def]>> + rw[] + , + qabbrev_tac‘s1 = p_1’>> qabbrev_tac‘s2 = p_2’>> + qabbrev_tac‘r1 = p_1' ’>> qabbrev_tac‘t3 = p_2' ’>> + qabbrev_tac‘e1 = p_2'' ’>> + qabbrev_tac‘t4 = p_1'' ’>> + qabbrev_tac‘e3 = p_2'³'’>> + qabbrev_tac‘t5 = p_2'''' ’>> + qabbrev_tac‘e4 = SP_csub sp e e3’>> + qabbrev_tac‘e5 = SP_csub sp e e2’>> + ‘e4 ∈ sp.Challenges.carrier’ by metis_tac[SP_csub_Lemma]>> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def)>> + ‘((sp.SimulatorMapInverse s2 w e4 t5,t4),e3) = ((r,t),e2)’by metis_tac[]>> + ‘r = sp.SimulatorMapInverse s2 w e4 t5 ∧ t = t4 ∧ e2 = e3’by rw[pairTheory.PAIR]>> + ‘e4 = SP_csub sp e e3’by metis_tac[]>> + ‘e3 = e2’by rw[]>> + ‘e5 = SP_csub sp e e2’by rw[]>> + ‘e4 = e5’by rw[]>> + ‘sp.SimulatorMap s2 w e5 (sp.SimulatorMapInverse s2 w e4 t5) = t5’by metis_tac[HonestVerifierZeroKnowledge_SP_def]>> + ‘((t,e2),sp.SimulatorMap s2 w e5 r) = ((t4,e3),t5)’suffices_by rw[]>> + ‘t = t4 ∧ e2 = e3’by rw[]>> + ‘r = sp.SimulatorMapInverse s2 w e4 t5 ∧ e4 = e5’by rw[]>> + ‘sp.SimulatorMap s2 w e5 r = t5’by rw[]>> + rw[] + ])); + + +(* Equality composition preserves well-formedness - if the input protocol is + well-formed, so is its equality composition *) +val WellFormed_SP_eq_thm = store_thm( + "WellFormed_SP_eq_thm", + “WellFormed_SP sp ⇒ WellFormed_SP (SP_eq sp)”, + rw[] >> + simp[WellFormed_SP_def, SP_eq_def, pairTheory.FORALL_PROD] >> + rpt (strip_tac >> + gvs[WellFormed_SP_def, SP_eq_def]) >> + rpt (rpt (pairarg_tac >> gvs[]) >> + fs[Gcross_def, pairTheory.PAIR] >> + gs[WellFormed_SP_def, pairTheory.PAIR] >> + metis_tac[pairTheory.PAIR])); + + +(* Stronger well-formedness conditions preserve under equality composition - when +the input protocol satisfies the special equality well-formedness properties, +its equality composition does too. *) +val Eq_WellFormed_SP_eq_thm = store_thm( + "Eq_WellFormed_SP_eq_thm", + “Eq_WellFormed_SP sp ⇒ Eq_WellFormed_SP (SP_eq sp)”, + rw[] >> + simp[Eq_WellFormed_SP_def, WellFormed_SP_def, SP_eq_def, pairTheory.FORALL_PROD] >> + rpt (strip_tac >> + gvs[Eq_WellFormed_SP_def, WellFormed_SP_def, SP_eq_def]) >> + rpt (rpt (pairarg_tac >> gvs[]) >> + fs[Gcross_def, pairTheory.PAIR] >> + gs[Eq_WellFormed_SP_def, WellFormed_SP_def, pairTheory.PAIR] >> + metis_tac[pairTheory.PAIR])); + + +(* Simulator correctness preserves under equality composition - if the input protocol +has correct simulation and satisfies equality well-formedness, then its equality +composition also produces valid simulated proofs. *) +val SimulatorCorrectness_SP_eq_thm = store_thm( + "SimulatorCorrectness_SP_eq_thm", + “SimulatorCorrectness_SP sp ∧ Eq_WellFormed_SP sp ⇒ SimulatorCorrectness_SP (SP_eq sp)”, + rw[] >> + simp[SimulatorCorrectness_SP_def, SP_eq_def, Eq_WellFormed_SP_def, WellFormed_SP_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + rpt (pairarg_tac >> gvs[]) >> + metis_tac[SimulatorCorrectness_SP_def, SP_eq_def, Eq_WellFormed_SP_def, WellFormed_SP_def]); + + +(* Completeness preserves under equality composition - if the input protocol is complete +and satisfies equality well-formedness conditions, then an honest prover can still +convince an honest verifier in the equality composition. *) +val Complete_SP_eq_thm = store_thm( + "Complete_SP_eq_thm", + “Complete_SP sp ∧ Eq_WellFormed_SP sp ⇒ Complete_SP (SP_eq sp)”, + rw[] >> + simp[Complete_SP_def, SP_eq_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >- + metis_tac[Complete_SP_def, SP_eq_def, SimulatorCorrectness_SP_def, Eq_WellFormed_SP_def, WellFormed_SP_def] >> + ‘WellFormed_SP sp’ by metis_tac[Eq_WellFormed_SP_def] >> + drule_then assume_tac (cj 2 (iffLR Eq_WellFormed_SP_def)) >> + ‘∀s1 s2 w r c1 c2 e. + s1 ∈ sp.Statements ∧ s2 ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ c1 ∈ sp.Commitments ∧ c2 ∈ sp.Commitments ∧ + e ∈ sp.Challenges.carrier ⇒ + sp.Prover_1 s1 w r c1 e = sp.Prover_1 s2 w r c2 e’ by metis_tac[] >> + qabbrev_tac‘c1 = sp.Prover_0 p_1 w r’ >> + qabbrev_tac‘c2 = sp.Prover_0 p_2 w r’ >> + ‘c1 ∈ sp.Commitments ∧ c2 ∈ sp.Commitments’ by metis_tac[Eq_WellFormed_SP_def, WellFormed_SP_def] >> + qabbrev_tac‘t1 = sp.Prover_1 p_1 w r c1 e’ >> + qabbrev_tac‘t2 = sp.Prover_1 p_2 w r c2 e’ >> + ‘sp.Prover_1 p_1 w r c1 e = sp.Prover_1 p_2 w r c2 e’ by metis_tac[Eq_WellFormed_SP_def, WellFormed_SP_def] >> + metis_tac[Complete_SP_def]); + + +(* Special soundness preserves under equality composition - if the input protocol is +specially sound and equality well-formed, its equality composition maintains the +ability to extract witnesses from accepting transcripts. *) +val SpecialSoundness_SP_eq_thm = store_thm( + "SpecialSoundness_SP_eq_thm", + “SpecialSoundness_SP sp ∧ Eq_WellFormed_SP sp ⇒ SpecialSoundness_SP (SP_eq sp)”, + simp[SpecialSoundness_SP_def, SP_eq_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + rpt (metis_tac[SpecialSoundness_SP_def, SP_eq_def, Eq_WellFormed_SP_def])); + + +(* Honest-verifier zero-knowledge preserves under equality composition - given a complete, +HVZK protocol satisfying equality well-formedness, its equality composition remains +zero-knowledge to honest verifiers. *) +val HonestVerifierZeroKnowledge_SP_eq_thm = store_thm( + "HonestVerifierZeroKnowledge_SP_eq_thm", + “HonestVerifierZeroKnowledge_SP sp ∧ Eq_WellFormed_SP sp ∧ Complete_SP sp ⇒ HonestVerifierZeroKnowledge_SP (SP_eq sp)”, + rw[] >> + simp[HonestVerifierZeroKnowledge_SP_def, SP_eq_def, pairTheory.FORALL_PROD] >> + rpt strip_tac >> + rpt (pairarg_tac >> gvs[]) >> + rw[] >> + rpt ( + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def) >> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘r’, ‘e’] mp_tac >> + ‘sp.Simulator p_1 (sp.SimulatorMap p_1 w e r) e = (p_1, sp.Prover_0 p_1 w r, e, sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e)’ by metis_tac[] >> + ‘sp.Simulator p_1 (sp.SimulatorMap p_1 w e r) e = (s1',c1',e1',t1')’ by metis_tac[] >> + ‘WellFormed_SP sp’ by metis_tac[Eq_WellFormed_SP_def] >> + ‘∀s t e s' c' e' t'. s ∈ sp.Statements ∧ t ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier ∧ sp.Simulator s t e = (s', c', e', t') ⇒ s' = s’ by metis_tac[WellFormed_SP_def] >> + ‘p_1 ∈ sp.Statements ∧ sp.SimulatorMap p_1 w e r ∈ sp.Responses ∧ e ∈ sp.Challenges.carrier’ by metis_tac[WellFormed_SP_def] >> + metis_tac[]) >- + (drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def) >> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘r’, ‘e’] mp_tac >> + ‘p_1 ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ∧ sp.Relation p_1 w’ by metis_tac[] >> + ‘sp.Simulator p_1 (sp.SimulatorMap p_1 w e r) e = (p_1, sp.Prover_0 p_1 w r, e, sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e)’ by metis_tac[] >> + ‘(s1',c1',e1',t1') = (p_1,sp.Prover_0 p_1 w r,e,sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e)’ by metis_tac[] >> + ‘FST(SND(s1',c1',e1',t1')) = FST(SND(p_1,sp.Prover_0 p_1 w r,e,sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e))’ by metis_tac[pairTheory.PAIR] >> + rw[]) >- + (‘WellFormed_SP sp’ by metis_tac[Eq_WellFormed_SP_def] >> + ‘∀s1 s2 w r e. s1 ∈ sp.Statements ∧ s2 ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ + r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ⇒ + sp.SimulatorMap s1 w e r = sp.SimulatorMap s2 w e r’ + by metis_tac[Eq_WellFormed_SP_def] >> + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def) >> + first_assum$ qspecl_then [‘p_2’, ‘w’, ‘r’, ‘e’] mp_tac >> + ‘sp.SimulatorMap p_1 w e r = sp.SimulatorMap p_2 w e r’ by metis_tac[] >> + ‘sp.Simulator p_2 (sp.SimulatorMap p_2 w e r) e = (p_2, sp.Prover_0 p_2 w r, e, sp.Prover_1 p_2 w r (sp.Prover_0 p_2 w r) e)’ by metis_tac[] >> + gvs[]) >> + rpt ( + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def) >> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘r’, ‘e’] mp_tac >> + ‘p_1 ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ∧ sp.Relation p_1 w’ by metis_tac[] >> + ‘sp.Simulator p_1 (sp.SimulatorMap p_1 w e r) e = (p_1, sp.Prover_0 p_1 w r, e, sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e)’ by metis_tac[] >> + ‘(s1',c1',e1',t1') = (p_1,sp.Prover_0 p_1 w r,e,sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e)’ by metis_tac[] >> + ‘FST(SND(s1',c1',e1',t1')) = FST(SND(p_1,sp.Prover_0 p_1 w r,e,sp.Prover_1 p_1 w r (sp.Prover_0 p_1 w r) e))’ by metis_tac[pairTheory.PAIR] >> + rw[]) >> + rpt ( + drule_then assume_tac (iffLR HonestVerifierZeroKnowledge_SP_def) >> + first_assum$ qspecl_then [‘p_1’, ‘w’, ‘r’, ‘e’] mp_tac >> + ‘p_1 ∈ sp.Statements ∧ w ∈ sp.Witnesses ∧ r ∈ sp.RandomCoins ∧ e ∈ sp.Challenges.carrier ∧ sp.Relation p_1 w’ by metis_tac[] >> + ‘let c = sp.Prover_0 p_1 w r; spm = sp.SimulatorMap p_1 w e; spmi = sp.SimulatorMapInverse p_1 w e in spmi(spm r) = r’ by metis_tac[] >> + metis_tac[])); + + +val _ = export_theory(); diff --git a/examples/algebra/README.md b/examples/algebra/README.md index f631984044..882bfc86d0 100644 --- a/examples/algebra/README.md +++ b/examples/algebra/README.md @@ -4,54 +4,19 @@ These theories of abstract algebra form a useful collection of libraries, to support in particular the mechanisation of AKS algorithm. +This library is being re-structured, +see comments in pull requests [#1224](https://github.com/HOL-Theorem-Prover/HOL/pull/1224) and [#1235](https://github.com/HOL-Theorem-Prover/HOL/pull/1235). + ## Helper Theories * __lib__ is the extension of existing HOL libraries. - * `lib/helperNum`, more theorems on `arithmeticTheory`, `dividesTheory`, and `gcdTheory`. - * `lib/helperSet`, more theorems about sets and mappings, extends `pred_setTheory`. - * `lib/helperList`, more theorems about list manipulations, extends `listTheory`. - * `lib/helperFunction`, theorems on function equivalence. - * `lib/binomial`, properties of Pascal's triangle, for binomial coefficients. - * `lib/triangle`, properties of Leibniz's Denominator triangle, for consecutive LCM. - * `lib/Euler`, properties of number-theoretic sets, and Euler's φ function. - * `lib/Gauss`, more theorems about coprimes, φ function, and Gauss' Little Theorem. - * `lib/primes`, properties of two-factors, and a primality test. - * `lib/primePower`, properties of prime powers and divisors, study for consecutive LCM. - * `lib/logPower`, properties of perfect power, power free, and upper logarithm. - * `lib/sublist`, order-preserving sublist and properties. + * Most scripts are now incorporated into HOL4 theories, the rest are moved to `src/algebra/base`. ## Basic Theories * __monoid__ is an algebraic structure with a useful binary operation. - * `monoid/monoid`, properties of a monoid with identity. - * `monoid/monoidInstances`, examples of monoids, in particular modulo numbers. - * `monoid/monoidOrder`, order of element in a monoid and its properties. - * `monoid/monoidMap`, homomorphism and isomorphism between monoids. - * `monoid/submonoid`, properties of submonoids of a monoid. - * __group__ is a monoid with all its elements invertible. - * `group/group`, properties of group, as an invertible monoid. - * `group/groupInstances`, examples of groups, in particular prime modulo systems. - * `group/groupOrder`, extends monoidOrder to groupOrder. - * `group/groupCyclic`, theorems on cyclic group and generators. - * `group/subgroup`, theorems on subgroups, including cosets. - * `group/quotientGroup`, theorems of a group partitioned by a normal subgroup. - * `group/groupMap`, homomorphism and isomorphism between groups. - * `group/congruences`, application to number theory, mainly Fermat's Little Theorem. - * `group/corres`, the Correspondence Theorem for group theory. - * `group/finiteGroup`, properties of a finite group. - * `group/groupAction`, action of a group on a target. - * __ring__ is an algebraic structure with two related binary operations. - * `ring/ring`, properties of a ring, made up of a group and a monoid. - * `ring/ringInstances`, examples of rings, in particular modulo arithmetic systems. - * `ring/ringUnit`, properties of invertible elements in a ring (they form a group). - * `ring/ringDivides`, theorems of division for invertible elements in a ring. - * `ring/ringIdeal`, the concept of a well-behaved sub-structure, like the normal subgroup. - * `ring/quotientRing`, theorems of a ring partitioned by an ideal. - * `ring/ringMap`, homomorphism and isomorphism between rings. - * `ring/ringInteger`, the similarity of the multiples of ring unity and the integers. - * `ring/ringBinomial`, the binomial theorem in terms of ring integers. - * `ring/integralDomain`, properties of integral domain, a non-trivial ring without zero divisors. - * `ring/integralDomainInstances`, examples of integral domains, e.g. arithmetic in modulo prime. + * These theories have been relocated to `src/algebra/construction`. + * The `ring` theory here is a port from HOL-light, as a place-holder. * __field__ is a nontrivial ring with all nonzero elements invertible. * `field/field`, properties of a field, as an invertible (zero excluded) ring. diff --git a/examples/algebra/field/files.txt b/examples/algebra/field/files.txt deleted file mode 100644 index aa4deba969..0000000000 --- a/examples/algebra/field/files.txt +++ /dev/null @@ -1,62 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Field Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2014 *) -(* ------------------------------------------------------------------------- *) - -0 field -- field definition and basic properties. -* monoid -* group -* subgroup -* ring -* ringUnit -* ringDivides -* integralDomain - -1 fieldUnit -- units in a field, all nonzero elements of its multiplicative monoid. -* ring -* ringUnit -* ringDivides -* 0 field - -1 fieldOrder -- multiplicative order of elements in a field. Sets of field elements of the same order. -* Gauss -* groupOrder -* groupCyclic -* ring -* integralDomain -* 0 field - -1 fieldProduct -- product of a set of field elements, and its manipulation. -* binomial -* subgroup -* groupProduct -* ring -* 0 field - -1 fieldIdeal -- ideals in a field. Quotient ring is a field iff the ideal is maximal. -* ringUnit -* ringIdeal -* quotientRing -* 0 field - -2 fieldMap, maps between fields: homomorphism and isomorphism. -* ringMap -* ringDivides -* ringUnit -* integralDomain -* 0 field -* 1 fieldOrder - -3 fieldInstances -- instances of field: (ZN p), integers within a prime modulus p. -* groupOrder -* ringInstances -* 0 field -* 2 fieldMap - -3 fieldBinomial -- binomial coefficients and binomial expansion in a field, the Freshmen theorem. -* binomial -* ringBinomial -* 0 field -* 2 fieldMap diff --git a/examples/algebra/finitefield/files.txt b/examples/algebra/finitefield/files.txt deleted file mode 100644 index ef67d8f5b1..0000000000 --- a/examples/algebra/finitefield/files.txt +++ /dev/null @@ -1,84 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Finite Field Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2014 *) -(* ------------------------------------------------------------------------- *) - -0 ffBasic -- basic properties of a finite field: no vector space. -* polyFieldModuloTheory -* fieldTheory -* fieldInstancesTheory - -1 ffInstances -- instances of finite field: GF_2, GF_4, GF_p for prime p. -* 0 ffBasic - -1 ffAdvanced -- advanced properties of a finite field: as a vector space over its subfield. -* VectorSpace -* FiniteVSpace -* 0 ffBasic - -2 ffPoly -- subring polynomials, common polynomials, roots of subfield polynomials. -* polyMultiplicity -* 0 ffBasic -* 1 ffAdvanced - -3 ffCyclo -- cyclotomic polynomials, the order of its roots. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly - -3 ffMaster -- master polynomials, relationship with irreducible polynomials of a subfield. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly - -4 ffUnity -- unity polynomials, roots of unity and the number of elements of each field order. -* 0 ffBasic -* 1 ffAdvanced -* 3 ffCyclo -* 3 ffMaster - -5 ffMinimal -- minimal polynomials, its existence by linear independence, and its properties. -* FiniteVSpace -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly -* 4 ffUnity - -6 ffConjugate -- conjugates of field elements, their order and product of conjugate factors. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly -* 3 ffCyclo -* 3 ffMaster -* 4 ffUnity -* 5 ffMinimal - -7 ffExist -- classification of finite fields: existence and uniqueness. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly -* 3 ffCyclo -* 3 ffMaster -* 4 ffUnity -* 5 ffMinimal -* 6 ffConjugate - -8 ffExtend -- field extension by isomorphic subfield. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly -* 5 ffMinimal -* 6 ffConjugate -* 7 ffExist - -9 ffSplit -- splitting field of a field polynomial, and finite field as a splitting field of its special master polynomial. -* 0 ffBasic -* 1 ffAdvanced -* 2 ffPoly -* 3 ffCyclo -* 4 ffUnity -* 6 ffConjugate -* 7 ffExist -* 8 ffExtend diff --git a/examples/algebra/linear/files.txt b/examples/algebra/linear/files.txt deleted file mode 100644 index aa6c845905..0000000000 --- a/examples/algebra/linear/files.txt +++ /dev/null @@ -1,21 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Linear Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2014 *) -(* ------------------------------------------------------------------------- *) - -0 VectorSpace -- vector space definition and basic properties. -* fieldTheory - -1 SpanSpace -- basis, vector sum, and spanning space. -* 0 VectorSpace - -2 LinearIndep -- change of basis and linear independence. -* 0 VectorSpace -* 1 SpanSpace - -3 FiniteVSpace -- finite Vector Space, subspaces and dimension. -* 0 VectorSpace -* 1 SpanSpace -* 2 LinearIndep diff --git a/examples/algebra/polynomial/files.txt b/examples/algebra/polynomial/files.txt deleted file mode 100644 index 80b4813c4b..0000000000 --- a/examples/algebra/polynomial/files.txt +++ /dev/null @@ -1,209 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Polynomial Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2014 *) -(* ------------------------------------------------------------------------- *) - -0 polynomial -- polynomial definition and basic properties. -* monoid -* group -* ring -* integralDomain -* field - -1 polyWeak -- raw polynomials: not normalised, may have leading zero. -* 0 polynomial - -2 polyRing -- polynomials with coefficients from a ring. -* 0 polynomial -* 1 polyWeak - -3 polyField -- polynomials with coefficients from a field. -* 0 polynomial -* 1 polyWeak -* 2 polyRing - -3 polyDivision -- polynomial Division with quotient and remainder. -* 0 polynomial -* 1 polyWeak -* 2 polyRing - -4 polyModuloRing -- polynomial quotient ring by a modulus. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision - -4 polyMonic -- monic Polynomials, those with leading coefficient equals one. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision - -5 polyFieldDivision -- division of field polynomials, any nonzero polynomial can be a divisor. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic - -5 polyRoot -- polynomial factors and roots. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic - -6 polyEval -- polynomial evaluation, acting as a function. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 4 polyMonic -* 5 polyRoot - -6 polyDerivative -- formal derivative for polynomials, symbolic term by term. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 4 polyMonic -* 5 polyRoot - -7 polyBinomial -- polynomial binomial expansion formula. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyRoot -* 6 polyEval - -7 polyDivides -- polynomial divisibility, relation to vanishing of remainder. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyRoot -* 5 polyFieldDivision -* 6 polyEval - -8 polyProduct -- product of polynomials, properties, evaluation, and divisibility. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyFieldDivision -* 5 polyRoot -* 6 polyEval -* 7 polyDivides - -8 polyIrreducible -- irreducible polynomials, those without proper factors. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyFieldDivision -* 5 polyRoot; -* 7 polyDivides - -8 polyRingModulo -- ring polynomial remainder congruences. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyDivision -* 3 polyField -* 4 polyModuloRing -* 5 polyFieldDivision -* 6 polyEval -* 7 polyDivides -* 7 polyBinomial - -9 polyCyclic -- properties of cyclic polynomial: quotient of (unity n) by (unity 1). -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyDivision -* 4 polyMonic -* 5 polyRoot -* 6 polyEval -* 7 polyDivides -* 8 polyIrreducible - -9 polyFieldModulo -- field polynomial remainder congruences. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 4 polyModuloRing -* 5 polyFieldDivision -* 5 polyRoot -* 6 polyEval -* 7 polyDivides -* 7 polyBinomial -* 8 polyRingModulo -* 8 polyIrreducible -* 8 polyProduct - -10 polyGCD -- greatest common divisor of polynomials. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyRoot -* 5 polyFieldDivision -* 6 polyEval -* 6 polyDerivative -* 7 polyDivides -* 8 polyProduct -* 8 polyIrreducible -* 9 polyFieldModulo - -11 polyMultiplicity -- multiplicity of polynomial factors and roots. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 5 polyRoot -* 6 polyEval -* 6 polyDerivative -* 7 polyBinomial -* 7 polyDivides -* 8 polyProduct -* 8 polyIrreducible -* 10 polyGCD - -12 polyMap -- maps between polynomials, under homomorphism or isomorphism of their coefficient rings or fields. -* 0 polynomial -* 1 polyWeak -* 2 polyRing -* 3 polyField -* 3 polyDivision -* 4 polyMonic -* 4 polyModuloRing -* 5 polyRoot -* 5 polyFieldDivision -* 6 polyEval -* 7 polyDivides -* 8 polyProduct -* 8 polyIrreducible -* 8 polyRingModulo -* 9 polyFieldModulo -* 11 polyMultiplicity diff --git a/examples/arm/arm6-verification/armScript.sml b/examples/arm/arm6-verification/armScript.sml index 01d40a7797..484745400c 100644 --- a/examples/arm/arm6-verification/armScript.sml +++ b/examples/arm/arm6-verification/armScript.sml @@ -845,9 +845,6 @@ val ARM_SPEC_STATE = save_thm("ARM_SPEC_STATE", (* ......................................................................... *) -infix \\ -val op \\ = op THEN; - val UPDATE_LT_COMMUTES = store_thm("UPDATE_LT_COMMUTES", `!m a b c d. a <+ b ==> ((b =+ d) ((a =+ c) m) = (a =+ c) ((b =+ d) m))`, diff --git a/examples/arm/arm6-verification/io_onestepScript.sml b/examples/arm/arm6-verification/io_onestepScript.sml index e7cb4e3f07..86bbe37cd7 100644 --- a/examples/arm/arm6-verification/io_onestepScript.sml +++ b/examples/arm/arm6-verification/io_onestepScript.sml @@ -19,11 +19,7 @@ val _ = ParseExtras.temp_loose_equality(); (* vvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvv *) -infix \\ << >> - -val op \\ = op THEN; -val op << = op THENL; -val op >> = op THEN1; +val op >- = op THEN1; val bool_ss = bool_ss ++ boolSimps.LET_ss; @@ -115,7 +111,7 @@ val LEAST_THM = store_thm("LEAST_THM", `!n. (!m. m < n ==> ~P m) /\ P n ==> ($LEAST P = n)`, REPEAT STRIP_TAC \\ IMP_RES_TAC whileTheory.FULL_LEAST_INTRO - \\ Cases_on `$LEAST P = n` >> ASM_REWRITE_TAC [] + \\ Cases_on `$LEAST P = n` >- ASM_REWRITE_TAC [] \\ `$LEAST P < n` by DECIDE_TAC \\ PROVE_TAC []); @@ -123,7 +119,7 @@ val lem = prove( `(!t1 t2. t1 < t2 ==> imm t1 < imm t2) ==> (!m. m < x ==> imm (m + 1) <= imm x)`, RW_TAC arith_ss [] - \\ Cases_on `m + 1 = x` >> ASM_SIMP_TAC arith_ss [] + \\ Cases_on `m + 1 = x` >- ASM_SIMP_TAC arith_ss [] \\ `m + 1 < x` by DECIDE_TAC \\ RES_TAC \\ DECIDE_TAC); @@ -163,13 +159,13 @@ val LIST_EQ = prove( val EL_GENLIST = prove( `!f n x. x < n ==> (EL x (GENLIST f n) = f x)`, - Induct_on `n` >> SIMP_TAC arith_ss [] + Induct_on `n` >- SIMP_TAC arith_ss [] \\ REPEAT STRIP_TAC \\ REWRITE_TAC [GENLIST] \\ Cases_on `x < n` \\ POP_ASSUM (fn th => ASSUME_TAC (SUBS [(GSYM o SPECL [`f`,`n`]) LENGTH_GENLIST] th) \\ ASSUME_TAC th) - >> ASM_SIMP_TAC bool_ss [EL_SNOC] + >- ASM_SIMP_TAC bool_ss [EL_SNOC] \\ `x = LENGTH (GENLIST f n)` by FULL_SIMP_TAC arith_ss [LENGTH_GENLIST] \\ ASM_SIMP_TAC bool_ss [EL_LENGTH_SNOC] \\ REWRITE_TAC [LENGTH_GENLIST]); @@ -216,7 +212,7 @@ val lem4 = prove( \\ IMP_RES_TAC ((SIMP_RULE std_ss [] o INST [`P` |-> `\t. x < imm (t + 1)`] o SPEC `t`) LEAST_THM) \\ FULL_SIMP_TAC std_ss [FREE_IMMERSION_def] - \\ Cases_on `t` >> ASM_SIMP_TAC arith_ss [] + \\ Cases_on `t` >- ASM_SIMP_TAC arith_ss [] \\ PAT_X_ASSUM `!t. t < SUC n ==> p` (SPEC_THEN `n` (ASSUME_TAC o SIMP_RULE arith_ss [NOT_LESS])) \\ ASM_REWRITE_TAC [ADD1]); @@ -228,11 +224,11 @@ val PACK_SERIALIZE = store_thm("PACK_SERIALIZE", (!strm. SERIALIZE imm (PACK imm strm) = strm)`, RW_TAC bool_ss [PACKED_STRM_def] \\ REWRITE_TAC [FUN_EQ_THM,LIST_EQ] - << [ + THENL [ RW_TAC bool_ss [PACK_def,LENGTH_GENLIST,EL_GENLIST] \\ SIMP_TAC arith_ss [SERIALIZE_def,IMM_START_def,IMM_RET_def] \\ Cases_on `x` - << [ + THENL [ FULL_SIMP_TAC arith_ss [FREE_IMMERSION_def,IMM_LEN_def, (SIMP_RULE arith_ss [NOT_LESS] o INST [`P` |-> `\t. n < imm (t + 1)`] o SPEC `0`) LEAST_THM], @@ -376,7 +372,7 @@ val PUIMMERSION_MONO_LEMMA2 = prove( REPEAT STRIP_TAC \\ IMP_RES_TAC PUIMMERSION_MONO_LEMMA \\ Induct_on `p` - >> ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,PUIMMERSION_MONO_LEMMA] + >- ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,PUIMMERSION_MONO_LEMMA] \\ FULL_SIMP_TAC bool_ss [SUC_COMM_LEMMA,LESS_IMP_LESS_ADD,ADD_COMM,PUIMMERSION_def]); @@ -459,7 +455,7 @@ val PTCON_THM = prove( Pstate_out_state (f t x))) ==> PTCON f strm`, RW_TAC bool_ss [PTCON_def,IS_PMAP_INIT_def] - << [ + THENL [ `IS_PMAP f` by PROVE_TAC [IS_PMAP_def] \\ IMP_RES_TAC PMAP2 \\ ASM_SIMP_TAC bool_ss [], FULL_SIMP_TAC bool_ss [PMAP_def] @@ -496,7 +492,7 @@ val PTCON_IMMERSION_LEMMA = prove( Pstate_out_state (f (imm x t) x)))`, RW_TAC bool_ss [IS_PMAP_INIT_def,IMMERSION,PTCON_IMMERSION] \\ EQ_TAC \\ REPEAT STRIP_TAC - << [ + THENL [ `IS_PMAP f` by PROVE_TAC [IS_PMAP_def] \\ PAT_X_ASSUM `!t1:num t2 x. P` (fn th => ASSUME_TAC (GSYM (SPECL [`0`,`t`] th))) @@ -553,7 +549,7 @@ val TC_STRM_ONE_STEP_THM = prove( `!strm. (!t i. i IN strm ==> (ADVANCE t i)) IN strm = (!i. i IN strm ==> (ADVANCE 1 i) IN strm)`, STRIP_TAC \\ EQ_TAC \\ REPEAT STRIP_TAC - >> ASM_SIMP_TAC bool_ss [] + >- ASM_SIMP_TAC bool_ss [] \\ Induct_on `t` \\ ASM_SIMP_TAC bool_ss [ADVANCE_ZERO,ADVANCE_ONE] ); @@ -604,13 +600,13 @@ val PTCON_IMMERSION_ONE_STEP_THM = prove( \\ EQ_TAC \\ STRIP_TAC \\ ASM_SIMP_TAC bool_ss [PTCON_IMMERSION_LEMMA2] \\ IMP_RES_TAC PTCON_IMMERSION_LEMMA - >> RW_TAC bool_ss [] + >- RW_TAC bool_ss [] \\ NTAC 3 (POP_ASSUM (K ALL_TAC)) \\ POP_ASSUM (fn th => REWRITE_TAC [th]) \\ Induct - >> FULL_SIMP_TAC bool_ss [PUNIFORM_def,PUIMMERSION_def,ADVANCE_ZERO] + >- FULL_SIMP_TAC bool_ss [PUNIFORM_def,PUIMMERSION_def,ADVANCE_ZERO] \\ REPEAT STRIP_TAC - << [ + THENL [ `imm x (SUC t) = imm (Pstate_out_state (f (imm x t) x)) 1 + imm x t` by PROVE_TAC [SPLIT_ITER_LEMMA2] \\ `IS_PMAP f` by PROVE_TAC [IS_PMAP_THM] @@ -700,7 +696,7 @@ val lem = prove( IS_PMAP f /\ PUNIFORM imm f /\ PTCON_IMMERSION f imm strm ==> !t x. x.inp IN strm ==> (imm (Pstate_out_state (f 0 x)) t = imm x t)`, RW_TAC bool_ss [PUNIFORM_def,PUIMMERSION_def] - \\ Induct_on `t` >> ASM_REWRITE_TAC [] + \\ Induct_on `t` >- ASM_REWRITE_TAC [] \\ FULL_SIMP_TAC bool_ss [PTCON_IMMERSION] \\ PAT_X_ASSUM `!t1 t2 x. P` (SPECL_THEN [`t`,`0`,`x`] IMP_RES_TAC) \\ PAT_X_ASSUM `!x. imm x 0 = 0` @@ -829,7 +825,7 @@ val PONE_STEP_THM = prove( \\ IMP_RES_TAC ONE_STEP_LEMMA \\ NTAC 2 (POP_ASSUM (K ALL_TAC)) \\ RW_TAC bool_ss [PCORRECT_def] - << [ + THENL [ IMP_RES_TAC PUNIFORM_IMP_IMMERSION, PAT_X_ASSUM `IS_PMAP impl` (K ALL_TAC), SPECL_THEN [`sstrm`,`istrm`,`spec`,`impl`,`imm`,`abs`,`ismpl`,`f`] @@ -854,7 +850,7 @@ val PONE_STEP_THM = prove( \\ PAT_X_ASSUM `a ==> b` MATCH_MP_TAC \\ REPEAT STRIP_TAC ] - \\ Induct_on `t` >> ASM_SIMP_TAC bool_ss [] + \\ Induct_on `t` >- ASM_SIMP_TAC bool_ss [] \\ PAT_X_ASSUM `!x. x.inp IN istrm ==> !t. P` IMP_RES_TAC \\ POP_ASSUM (fn th => REWRITE_TAC [th]) \\ PAT_X_ASSUM `PTCON_IMMERSION impl imm istrm` @@ -1267,7 +1263,7 @@ val DATA_ABSTRACTION_I = store_thm("DATA_ABSTRACTION_I", (!a. ?b. abs (h b) = a))`, RW_TAC (srw_ss()++boolSimps.LET_ss) [IS_IMAP_INIT_def,IMAP_def, DATA_ABSTRACTION_def,RANGE_def,IMAGE_DEF,SURJ_DEF,IN_UNIV,GSPECIFICATION] - \\ Tactical.REVERSE EQ_TAC >> METIS_TAC [lem,lem2] + \\ Tactical.REVERSE EQ_TAC >- METIS_TAC [lem,lem2] \\ REPEAT STRIP_TAC \\ PAT_X_ASSUM `!x. (f 0 x).state = y` (SPEC_THEN `<|state := a; inp := i|>` (ASSUME_TAC o SIMP_RULE (srw_ss()) [] o SYM)) @@ -1300,7 +1296,7 @@ val CORRECT = prove( RW_TAC (srw_ss()++boolSimps.LET_ss) [CORRECT_def] \\ EQ_TAC \\ REPEAT STRIP_TAC \\ FULL_SIMP_TAC (srw_ss()) [] - << [ + THENL [ PAT_X_ASSUM `!x. P` (SPEC_THEN `<| state := a; inp := i |>` (IMP_RES_TAC o SIMP_RULE (srw_ss()) [])), PAT_X_ASSUM `!x. P` (SPEC_THEN `<| state := a; inp := i |>` @@ -1353,7 +1349,7 @@ val STATE_FUNPOW_INIT4 = store_thm("STATE_FUNPOW_INIT4", \\ IMP_RES_TAC STATE_FUNPOW_INIT \\ SPEC_TAC (`t`,`t`) \\ Induct - >> ASM_SIMP_TAC (srw_ss()++boolSimps.LET_ss) [FUNPOW,PINIT_def,SINIT_def] + >- ASM_SIMP_TAC (srw_ss()++boolSimps.LET_ss) [FUNPOW,PINIT_def,SINIT_def] \\ FULL_SIMP_TAC (srw_ss() ++ boolSimps.LET_ss) [FUNPOW_SUC, PNEXT_def, SNEXT_def, SIMP_RULE (srw_ss()) [PNEXT_def] STATE_FUNPOW_INIT2, diff --git a/examples/arm/arm6-verification/lemmasScript.sml b/examples/arm/arm6-verification/lemmasScript.sml index a36989427d..916e2cc083 100644 --- a/examples/arm/arm6-verification/lemmasScript.sml +++ b/examples/arm/arm6-verification/lemmasScript.sml @@ -18,11 +18,6 @@ val _ = new_theory "lemmas"; (* ------------------------------------------------------------------------- *) -infix << >> - -val op << = op THENL; -val op >> = op THEN1; - val std_ss = std_ss ++ boolSimps.LET_ss; val arith_ss = arith_ss ++ boolSimps.LET_ss; @@ -392,15 +387,15 @@ val lem = prove( val SLICE_ROR_THM = store_thm("SLICE_ROR_THM", `!h l a. ((h '' l) a) #>> l = (h -- l) a`, REPEAT STRIP_TAC \\ Cases_on `l = 0` - >> ASM_REWRITE_TAC [WORD_SLICE_BITS_THM,SHIFT_ZERO] + >- ASM_REWRITE_TAC [WORD_SLICE_BITS_THM,SHIFT_ZERO] \\ Cases_on `a` \\ RW_TAC arith_ss [MIN_DEF,word_slice_n2w,word_bits_n2w,BITS_COMP_THM2, SLICE_THM,w2n_n2w] - << [ - Cases_on `h < l` >> ASM_SIMP_TAC arith_ss [BITS_ZERO,ZERO_SHIFT] + THENL [ + Cases_on `h < l` >- ASM_SIMP_TAC arith_ss [BITS_ZERO,ZERO_SHIFT] \\ `l <= dimindex (:'a) - 1` by DECIDE_TAC, Cases_on `dimindex (:'a) - 1 < l` - >> ASM_SIMP_TAC arith_ss [BITS_ZERO,ZERO_SHIFT]] + >- ASM_SIMP_TAC arith_ss [BITS_ZERO,ZERO_SHIFT]] \\ RW_TAC arith_ss [BITS_ZERO3,ADD1,lem,word_ror_n2w, ZERO_LT_TWOEXP,ONCE_REWRITE_RULE [MULT_COMM] MOD_EQ_0] \\ ASM_SIMP_TAC arith_ss [BITS_SLICE_THM2, diff --git a/examples/arm/arm6-verification/my_listScript.sml b/examples/arm/arm6-verification/my_listScript.sml index 5e3542160f..e53fb9d522 100644 --- a/examples/arm/arm6-verification/my_listScript.sml +++ b/examples/arm/arm6-verification/my_listScript.sml @@ -5,11 +5,7 @@ val _ = new_theory "my_list"; (* ------------------------------------------------------------------------- *) -infix \\ << >> - -val op \\ = op THEN; -val op << = op THENL; -val op >> = op THEN1; +val op >- = op THEN1; (* ------------------------------------------------------------------------- *) @@ -17,9 +13,9 @@ val LIST_EQ = store_thm("LIST_EQ", `!a b. (LENGTH a = LENGTH b) /\ (!x. x < LENGTH a ==> (EL x a = EL x b)) ==> (a = b)`, - Induct_on `b` >> SIMP_TAC list_ss [LENGTH_NIL] + Induct_on `b` >- SIMP_TAC list_ss [LENGTH_NIL] \\ Induct_on `a` \\ RW_TAC list_ss [] - >> POP_ASSUM (fn th => PROVE_TAC [(SIMP_RULE list_ss [] o SPEC `0`) th]) + >- POP_ASSUM (fn th => PROVE_TAC [(SIMP_RULE list_ss [] o SPEC `0`) th]) \\ POP_ASSUM (fn th => PROVE_TAC [(GEN_ALL o SIMP_RULE list_ss [] o SPEC `SUC x`) th])); @@ -30,14 +26,14 @@ val NULL_SNOC = store_thm("NULL_SNOC", val FILTER_ALL = store_thm("FILTER_ALL", `!l. (!n. n < LENGTH l ==> ~P (EL n l)) ==> (FILTER P l = [])`, Induct \\ RW_TAC list_ss [] - >> (EXISTS_TAC `0` \\ ASM_SIMP_TAC list_ss []) + >- (EXISTS_TAC `0` \\ ASM_SIMP_TAC list_ss []) \\ POP_ASSUM (fn th => PROVE_TAC [(GEN_ALL o SIMP_RULE list_ss [] o SPEC `SUC n`) th])); val FILTER_NONE = store_thm("FILTER_NONE", `!l. (!n. n < LENGTH l ==> P (EL n l)) ==> (FILTER P l = l)`, Induct \\ RW_TAC list_ss [] - >> POP_ASSUM (fn th => ASM_SIMP_TAC std_ss + >- POP_ASSUM (fn th => ASM_SIMP_TAC std_ss [(GEN_ALL o SIMP_RULE list_ss [] o SPEC `SUC n`) th]) \\ POP_ASSUM (STRIP_ASSUME_TAC o SIMP_RULE list_ss [] o SPEC `0`)); @@ -47,13 +43,13 @@ val MAP_GENLIST = store_thm("MAP_GENLIST", val EL_GENLIST = store_thm("EL_GENLIST", `!f n x. x < n ==> (EL x (GENLIST f n) = f x)`, - Induct_on `n` >> SIMP_TAC arith_ss [] + Induct_on `n` >- SIMP_TAC arith_ss [] \\ REPEAT STRIP_TAC \\ REWRITE_TAC [GENLIST] \\ Cases_on `x < n` \\ POP_ASSUM (fn th => ASSUME_TAC (SUBS [(GSYM o SPECL [`f`,`n`]) LENGTH_GENLIST] th) \\ ASSUME_TAC th) - >> ASM_SIMP_TAC bool_ss [EL_SNOC] + >- ASM_SIMP_TAC bool_ss [EL_SNOC] \\ `x = LENGTH (GENLIST f n)` by FULL_SIMP_TAC arith_ss [LENGTH_GENLIST] \\ ASM_SIMP_TAC bool_ss [EL_LENGTH_SNOC] \\ REWRITE_TAC [LENGTH_GENLIST]); @@ -99,10 +95,10 @@ val ZIP_APPEND = store_thm("ZIP_APPEND", `!a b c d. (LENGTH a = LENGTH b) /\ (LENGTH c = LENGTH d) ==> (ZIP (a,b) ++ ZIP (c,d) = ZIP (a ++ c,b ++ d))`, - Induct_on `b` >> SIMP_TAC list_ss [LENGTH_NIL] - \\ Induct_on `d` >> SIMP_TAC list_ss [LENGTH_NIL] - \\ Induct_on `a` >> SIMP_TAC list_ss [LENGTH_NIL] - \\ Induct_on `c` >> SIMP_TAC list_ss [LENGTH_NIL] + Induct_on `b` >- SIMP_TAC list_ss [LENGTH_NIL] + \\ Induct_on `d` >- SIMP_TAC list_ss [LENGTH_NIL] + \\ Induct_on `a` >- SIMP_TAC list_ss [LENGTH_NIL] + \\ Induct_on `c` >- SIMP_TAC list_ss [LENGTH_NIL] \\ MAP_EVERY X_GEN_TAC [`h1`,`h2`,`h3`,`h4`] \\ RW_TAC list_ss [] \\ `LENGTH (h1::c) = LENGTH (h3::d)` diff --git a/examples/arm/arm6-verification/onestepScript.sml b/examples/arm/arm6-verification/onestepScript.sml index 5d7f67b8e1..9ad6eda393 100644 --- a/examples/arm/arm6-verification/onestepScript.sml +++ b/examples/arm/arm6-verification/onestepScript.sml @@ -12,14 +12,6 @@ open combinTheory arithmeticTheory prim_recTheory pred_setTheory; val _ = new_theory "onestep"; -(* vvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvv *) - -infix \\ << >> - -val op \\ = op THEN; -val op << = op THENL; -val op >> = op THEN1; - (* ------------------------------------------------------------------------- *) val FUNPOW_THM = store_thm("FUNPOW_THM", @@ -158,7 +150,7 @@ val UIMMERSION_MONO_LEMMA2 = prove( REPEAT STRIP_TAC \\ IMP_RES_TAC UIMMERSION_MONO_LEMMA \\ Induct_on `p` - >> ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,UIMMERSION_MONO_LEMMA] + >- ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,UIMMERSION_MONO_LEMMA] \\ FULL_SIMP_TAC bool_ss [SUC_COMM_LEMMA,LESS_IMP_LESS_ADD,ADD_COMM,UIMMERSION_def]); @@ -190,7 +182,7 @@ val AUIMMERSION_MONO_LEMMA2 = prove( REPEAT STRIP_TAC \\ IMP_RES_TAC AUIMMERSION_MONO_LEMMA \\ Induct_on `p` - >> ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,AUIMMERSION_MONO_LEMMA] + >- ASM_REWRITE_TAC [SYM ONE,GSYM ADD1,AUIMMERSION_MONO_LEMMA] \\ FULL_SIMP_TAC bool_ss [SUC_COMM_LEMMA,LESS_IMP_LESS_ADD,ADD_COMM,AUIMMERSION_def]); @@ -254,7 +246,7 @@ val TC_THM = store_thm("TC_THM", `!f. TCON f <=> IS_IMAP f /\ (!t a. f 0 (f t a) = f t a)`, STRIP_TAC \\ EQ_TAC \\ REPEAT STRIP_TAC - << [ + THENL [ PROVE_TAC [TC_IMP_IMAP], FULL_SIMP_TAC bool_ss [TCON_def] \\ POP_ASSUM (fn th => REWRITE_TAC @@ -307,7 +299,7 @@ val TC_IMMERSION_LEMMA2 = store_thm("TC_IMMERSION_LEMMA2", (TCON_IMMERSION f imm = !t a. f 0 (f (imm a t) a) = f (imm a t) a)`, RW_TAC bool_ss [IS_IMAP_def,IMMERSION,TCON_IMMERSION_def] \\ EQ_TAC - \\ REPEAT STRIP_TAC << [ + \\ REPEAT STRIP_TAC THENL [ POP_ASSUM (fn th => ASSUME_TAC (SPECL [`0`,`t`] th)) \\ PAT_ASSUM `!a. x /\ y` (fn th => RULE_ASSUM_TAC (SIMP_RULE arith_ss [th])) @@ -350,7 +342,7 @@ val TC_IMMERSION_THM = store_thm("TC_IMMERSION_THM", \\ IMP_RES_TAC UIMMERSION_IS_IMMERSION_LEMMA \\ IMP_RES_TAC TC_IMMERSION_LEMMA \\ FULL_SIMP_TAC bool_ss [IS_IMAP_def,TC_THM] - \\ STRIP_TAC >> PROVE_TAC [] + \\ STRIP_TAC >- PROVE_TAC [] \\ IMP_RES_TAC TC_IMM_LEM \\ FULL_SIMP_TAC bool_ss [IMAP_def]); (*--------------------------------------------------------------------------- @@ -375,7 +367,7 @@ val TC_IMMERSION_ONE_STEP_THM = store_thm("TC_IMMERSION_ONE_STEP_THM", \\ IMP_RES_TAC UNIFORM_IMP_IMMERSION \\ EQ_TAC \\ STRIP_TAC \\ ASM_SIMP_TAC bool_ss [TC_IMMERSION_LEMMA,TC_IMMERSION_LEMMA2] - \\ Induct >> ASM_REWRITE_TAC [] + \\ Induct >- ASM_REWRITE_TAC [] \\ GEN_TAC \\ FULL_SIMP_TAC bool_ss [UNIFORM_def,IS_IMAP_def] \\ PAT_ASSUM `UIMMERSION imm f dur` (fn th => REWRITE_TAC [REWRITE_RULE [UIMMERSION_def] th] \\ ASSUME_TAC th) @@ -437,8 +429,8 @@ Proof \\ IMP_RES_TAC ONE_STEP_LEMMA \\ EQ_TAC \\ RW_TAC bool_ss [CORRECT_def] - >> IMP_RES_TAC UNIFORM_IMP_IMMERSION - \\ Induct_on `t` >> ASM_REWRITE_TAC [] + >- IMP_RES_TAC UNIFORM_IMP_IMMERSION + \\ Induct_on `t` >- ASM_REWRITE_TAC [] \\ PAT_ASSUM `!t a. imm a (SUC t) = imm (impl (imm a t) a) 1 + imm a t` (fn th => REWRITE_TAC [th]) \\ PAT_ASSUM `TCON_IMMERSION impl imm` @@ -477,10 +469,10 @@ Proof \\ IMP_RES_TAC ADJUNCT_IMP_UNIFORM \\ EQ_TAC \\ RW_TAC bool_ss [CORRECT_SUP_def] - << [ + THENL [ IMP_RES_TAC ADJUNCT_IMP_IMMERSIONS, IMP_RES_TAC ADJUNCT_IMP_IMMERSIONS, - Induct_on `r` >> ASM_REWRITE_TAC [] + Induct_on `r` >- ASM_REWRITE_TAC [] \\ IMP_RES_TAC ONE_STEP_SUP_LEMMA \\ POP_ASSUM (fn th => REWRITE_TAC [th]) \\ IMP_RES_TAC ONE_STEP_LEMMA diff --git a/examples/arm/v4/arm_evalScript.sml b/examples/arm/v4/arm_evalScript.sml index 215b74cdf8..a15d3d9b6a 100644 --- a/examples/arm/v4/arm_evalScript.sml +++ b/examples/arm/v4/arm_evalScript.sml @@ -20,11 +20,6 @@ val _ = ParseExtras.temp_loose_equality() (* ------------------------------------------------------------------------- *) -infix << >> - -val op << = op THENL; -val op >> = op THEN1; - val std_ss = std_ss ++ boolSimps.LET_ss; val arith_ss = arith_ss ++ boolSimps.LET_ss; @@ -395,7 +390,7 @@ val PSR_CONS = store_thm("PSR_CONS", (SET_IFMODE (w ' 7) (w ' 6) m (0xFFFFF20w && w))`, RW_TAC word_ss [SET_IFMODE_def,SET_NZCV_def,word_modify_def,n2w_def] \\ RW_TAC word_ss [word_bits_def] - << [ + THENL [ `(i = 31) \/ (i = 30) \/ (i = 29) \/ (i = 28) \/ (i < 28)` by DECIDE_TAC \\ ASM_SIMP_TAC arith_ss [], @@ -403,7 +398,7 @@ val PSR_CONS = store_thm("PSR_CONS", (7 < i /\ i < 28) \/ (i = 7) \/ (i = 6) \/ (i = 5) \/ (i < 5)` by DECIDE_TAC \\ ASM_SIMP_TAC (word_ss++ARITH_ss) [word_and_def] - << [ + THENL [ FULL_SIMP_TAC std_ss [LESS_THM] \\ FULL_SIMP_TAC arith_ss [] \\ EVAL_TAC, EVAL_TAC, @@ -783,10 +778,10 @@ val immediate_enc = store_thm("immediate_enc", [IMMEDIATE_def,addr_mode1_encode_def,msr_mode_encode_def] \\ (MATCH_MP_TAC (METIS_PROVE [] ``!a b c d x. (a = b) /\ (c = d) ==> (ROR a c x = ROR b d x)``) - \\ STRIP_TAC << [ALL_TAC, + \\ STRIP_TAC THENL [ALL_TAC, MATCH_MP_TAC (PROVE [] ``!a b. (a = b) ==> (2w:word8 * a = 2w * b)``)] \\ SRW_TAC word_frags [WORD_EQ] - << [Cases_on `i' < 12` \\ SRW_TAC word_frags [] + THENL [Cases_on `i' < 12` \\ SRW_TAC word_frags [] \\ Cases_on `i' < 8` \\ SRW_TAC word_frags [], Cases_on `i' < 4` \\ SRW_TAC word_frags []] \\ POP_ASSUM_LIST (ASSUME_TAC o hd) diff --git a/examples/diningcryptos/extra_stringScript.sml b/examples/diningcryptos/extra_stringScript.sml index ef396cf341..f8225e93b8 100644 --- a/examples/diningcryptos/extra_stringScript.sml +++ b/examples/diningcryptos/extra_stringScript.sml @@ -32,13 +32,6 @@ val _ = new_theory "extra_string"; (* Helpful proof tools *) (* ------------------------------------------------------------------------- *) -infixr 0 ++ << || THENC ORELSEC ORELSER ##; -infix 1 >>; - -val op ++ = op THEN; -val op << = op THENL; -val op >> = op THEN1; -val op || = op ORELSE; val REVERSE = Tactical.REVERSE; val Suff = PARSE_TAC SUFF_TAC; @@ -82,14 +75,14 @@ val STRCAT1 = store_thm val append_neq_lem = prove (``!m l. (~(l = [])) ==> (~(l ++ m = m))``, REPEAT STRIP_TAC - ++ `LENGTH m = LENGTH (l ++ m)` by RW_TAC std_ss [] - ++ Cases_on `m` >> FULL_SIMP_TAC list_ss [APPEND_NIL] - ++ Q.PAT_ASSUM `l ++ h::t = h::t` (K ALL_TAC) - ++ FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]); + >> `LENGTH m = LENGTH (l ++ m)` by RW_TAC std_ss [] + >> Cases_on `m` >- FULL_SIMP_TAC list_ss [APPEND_NIL] + >> Q.PAT_ASSUM `l ++ h::t = h::t` (K ALL_TAC) + >> FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]); val append_sing_eq_lem = prove (``!l l' x x'. (l ++ [x] = l' ++ [x']) = ((l = l') /\ (x = x'))``, - RW_TAC std_ss [GSYM SNOC_APPEND, SNOC_11] ++ DECIDE_TAC); + RW_TAC std_ss [GSYM SNOC_APPEND, SNOC_11] >> DECIDE_TAC); val IMPLODE_APPEND_EQ = store_thm ("IMPLODE_APPEND_EQ", @@ -97,12 +90,12 @@ val IMPLODE_APPEND_EQ = store_thm (IMPLODE (s1 ++ s2) = IMPLODE (s1' ++ s2')) = (s1 ++ s2 = s1' ++ s2')``, Induct - >> (Induct >> (RW_TAC std_ss [APPEND_NIL, IMPLODE_EQ_EMPTYSTRING, IMPLODE_EQNS] - ++ RW_TAC bool_ss [EQ_SYM_EQ]) - ++ FULL_SIMP_TAC std_ss [APPEND_NIL, IMPLODE_EQNS, IMPLODE_EQ_THM, EXPLODE_IMPLODE] - ++ RW_TAC bool_ss [EQ_SYM_EQ]) - ++ FULL_SIMP_TAC std_ss [APPEND, IMPLODE_EQNS, IMPLODE_EQ_THM, EXPLODE_IMPLODE] - ++ RW_TAC bool_ss [EQ_SYM_EQ]); + >- (Induct >- (RW_TAC std_ss [APPEND_NIL, IMPLODE_EQ_EMPTYSTRING, IMPLODE_EQNS] + >> RW_TAC bool_ss [EQ_SYM_EQ]) + >> FULL_SIMP_TAC std_ss [APPEND_NIL, IMPLODE_EQNS, IMPLODE_EQ_THM, EXPLODE_IMPLODE] + >> RW_TAC bool_ss [EQ_SYM_EQ]) + >> FULL_SIMP_TAC std_ss [APPEND, IMPLODE_EQNS, IMPLODE_EQ_THM, EXPLODE_IMPLODE] + >> RW_TAC bool_ss [EQ_SYM_EQ]); val STRCAT_NEQ = store_thm ("STRCAT_NEQ", @@ -111,46 +104,46 @@ val STRCAT_NEQ = store_thm (~ (IS_PREFIX (EXPLODE s1) (EXPLODE s1'))) /\ (~ (IS_PREFIX (EXPLODE s1') (EXPLODE s1))) ==> (!s2 s2'. ~(STRCAT s1 s2 = STRCAT s1' s2'))``, - STRIP_TAC ++ STRIP_TAC - ++ (ASSUME_TAC o Q.SPEC `s1`) IMPLODE_ONTO - ++ FULL_SIMP_TAC std_ss [] ++ POP_ASSUM (K ALL_TAC) - ++ (ASSUME_TAC o Q.SPEC `s1'`) IMPLODE_ONTO - ++ FULL_SIMP_TAC std_ss [] ++ POP_ASSUM (K ALL_TAC) - ++ RW_TAC std_ss [EXPLODE_IMPLODE, IMPLODE_EQ_EMPTYSTRING] - ++ (ASSUME_TAC o Q.SPEC `s2`) IMPLODE_ONTO - ++ FULL_SIMP_TAC std_ss [] ++ POP_ASSUM (K ALL_TAC) - ++ (ASSUME_TAC o Q.SPEC `s2'`) IMPLODE_ONTO - ++ FULL_SIMP_TAC std_ss [] ++ POP_ASSUM (K ALL_TAC) - ++ RW_TAC std_ss [EXPLODE_IMPLODE, Once STRCAT1, IMPLODE_APPEND_EQ] - ++ Q.ABBREV_TAC `foo = IMPLODE (STRCAT cs cs'')` - ++ RW_TAC std_ss [EXPLODE_IMPLODE, Once STRCAT1, IMPLODE_APPEND_EQ] - ++ Q.UNABBREV_TAC `foo` - ++ RW_TAC std_ss [EXPLODE_IMPLODE, IMPLODE_APPEND_EQ] - ++ SPOSE_NOT_THEN STRIP_ASSUME_TAC - ++ (ASSUME_TAC o Q.ISPECL [`LENGTH (cs:char list)`, + STRIP_TAC >> STRIP_TAC + >> (ASSUME_TAC o Q.SPEC `s1`) IMPLODE_ONTO + >> FULL_SIMP_TAC std_ss [] >> POP_ASSUM (K ALL_TAC) + >> (ASSUME_TAC o Q.SPEC `s1'`) IMPLODE_ONTO + >> FULL_SIMP_TAC std_ss [] >> POP_ASSUM (K ALL_TAC) + >> RW_TAC std_ss [EXPLODE_IMPLODE, IMPLODE_EQ_EMPTYSTRING] + >> (ASSUME_TAC o Q.SPEC `s2`) IMPLODE_ONTO + >> FULL_SIMP_TAC std_ss [] >> POP_ASSUM (K ALL_TAC) + >> (ASSUME_TAC o Q.SPEC `s2'`) IMPLODE_ONTO + >> FULL_SIMP_TAC std_ss [] >> POP_ASSUM (K ALL_TAC) + >> RW_TAC std_ss [EXPLODE_IMPLODE, Once STRCAT1, IMPLODE_APPEND_EQ] + >> Q.ABBREV_TAC `foo = IMPLODE (STRCAT cs cs'')` + >> RW_TAC std_ss [EXPLODE_IMPLODE, Once STRCAT1, IMPLODE_APPEND_EQ] + >> Q.UNABBREV_TAC `foo` + >> RW_TAC std_ss [EXPLODE_IMPLODE, IMPLODE_APPEND_EQ] + >> SPOSE_NOT_THEN STRIP_ASSUME_TAC + >> (ASSUME_TAC o Q.ISPECL [`LENGTH (cs:char list)`, `LENGTH (cs':char list)`]) LESS_EQ_CASES - ++ FULL_SIMP_TAC std_ss [] - >> ((ASSUME_TAC o Q.ISPEC `(cs''':char list)` o + >> FULL_SIMP_TAC std_ss [] + >- ((ASSUME_TAC o Q.ISPEC `(cs''':char list)` o UNDISCH o Q.ISPECL [`LENGTH (cs:char list)`,`(cs':char list)`]) FIRSTN_APPEND1 - ++ `FIRSTN (LENGTH (cs:char list)) (cs':char list) = cs` + >> `FIRSTN (LENGTH (cs:char list)) (cs':char list) = cs` by METIS_TAC [FIRSTN_APPEND1, LESS_EQ_REFL, FIRSTN_LENGTH_ID] - ++ FULL_SIMP_TAC std_ss [LESS_EQ_EXISTS] - ++ (ASSUME_TAC o UNDISCH o REWRITE_RULE [Once EQ_SYM_EQ, Once ADD_COMM] + >> FULL_SIMP_TAC std_ss [LESS_EQ_EXISTS] + >> (ASSUME_TAC o UNDISCH o REWRITE_RULE [Once EQ_SYM_EQ, Once ADD_COMM] o Q.ISPECL [`p:num`, `LENGTH (cs:char list)`,`(cs':char list)`] ) APPEND_FIRSTN_LASTN - ++ FULL_SIMP_TAC std_ss [IS_PREFIX_APPEND] - ++ METIS_TAC []) - ++ (ASSUME_TAC o Q.ISPEC `(cs'':char list)` o + >> FULL_SIMP_TAC std_ss [IS_PREFIX_APPEND] + >> METIS_TAC []) + >> (ASSUME_TAC o Q.ISPEC `(cs'':char list)` o UNDISCH o Q.ISPECL [`LENGTH (cs':char list)`,`(cs:char list)`]) FIRSTN_APPEND1 - ++ `FIRSTN (LENGTH (cs':char list)) (cs:char list) = cs'` + >> `FIRSTN (LENGTH (cs':char list)) (cs:char list) = cs'` by METIS_TAC [FIRSTN_APPEND1, LESS_EQ_REFL, FIRSTN_LENGTH_ID] - ++ FULL_SIMP_TAC std_ss [LESS_EQ_EXISTS] - ++ (ASSUME_TAC o UNDISCH o REWRITE_RULE [Once EQ_SYM_EQ, Once ADD_COMM] + >> FULL_SIMP_TAC std_ss [LESS_EQ_EXISTS] + >> (ASSUME_TAC o UNDISCH o REWRITE_RULE [Once EQ_SYM_EQ, Once ADD_COMM] o Q.ISPECL [`p:num`, `LENGTH (cs':char list)`,`(cs:char list)`]) APPEND_FIRSTN_LASTN - ++ FULL_SIMP_TAC std_ss [IS_PREFIX_APPEND] - ++ METIS_TAC []); + >> FULL_SIMP_TAC std_ss [IS_PREFIX_APPEND] + >> METIS_TAC []); (* ------------------------------------------------------------------------- *) @@ -160,91 +153,91 @@ val toString_inj = store_thm ("toString_inj", ``!n m. (toString n = toString m) = (n = m)``, completeInduct_on `n` - ++ REPEAT STRIP_TAC - ++ Cases_on `n` - >> (Cases_on `m` ++ RW_TAC arith_ss [] - ++ RW_TAC std_ss [toString_def] - ++ SRW_TAC [] [rec_toString_def] - ++ Cases_on `SUC n MOD 10 = 0` - >> (SRW_TAC [] [] - ++ Cases_on `SUC n DIV 10` - >> (METIS_TAC [SUC_NOT,MULT_EQ_0, ADD_0, ADD_COMM, + >> REPEAT STRIP_TAC + >> Cases_on `n` + >- (Cases_on `m` >> RW_TAC arith_ss [] + >> RW_TAC std_ss [toString_def] + >> SRW_TAC [] [rec_toString_def] + >> Cases_on `SUC n MOD 10 = 0` + >- (SRW_TAC [] [] + >> Cases_on `SUC n DIV 10` + >- (METIS_TAC [SUC_NOT,MULT_EQ_0, ADD_0, ADD_COMM, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM] - ++ MATCH_MP_TAC append_neq_lem - ++ RW_TAC list_ss []) - ++ Cases_on `rec_toString (SUC n DIV 10)` - >> (SRW_TAC [] [CHAR_EQ_THM] - ++ `ORD (CHR (48 + SUC n MOD 10)) = 48 + SUC n MOD 10` + >> RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM] + >> MATCH_MP_TAC append_neq_lem + >> RW_TAC list_ss []) + >> Cases_on `rec_toString (SUC n DIV 10)` + >- (SRW_TAC [] [CHAR_EQ_THM] + >> `ORD (CHR (48 + SUC n MOD 10)) = 48 + SUC n MOD 10` by (RW_TAC std_ss [GSYM ORD_CHR] - ++ MATCH_MP_TAC LESS_LESS_EQ_TRANS ++ Q.EXISTS_TAC `48 + 10:num` - ++ RW_TAC arith_ss [LT_ADD_LCANCEL, + >> MATCH_MP_TAC LESS_LESS_EQ_TRANS >> Q.EXISTS_TAC `48 + 10:num` + >> RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ RW_TAC arith_ss []) - ++ RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM]) - ++ Cases_on `m` - >> (RW_TAC std_ss [toString_def] - ++ RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM] - ++ Cases_on `SUC n' MOD 10 = 0` - >> (SRW_TAC [] [] - ++ Cases_on `SUC n' DIV 10` - >> (METIS_TAC [SUC_NOT,MULT_EQ_0, ADD_0, ADD_COMM, + >> RW_TAC arith_ss []) + >> RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM]) + >> Cases_on `m` + >- (RW_TAC std_ss [toString_def] + >> RW_TAC string_ss [rec_toString_def,IMPLODE_EQ_THM] + >> Cases_on `SUC n' MOD 10 = 0` + >- (SRW_TAC [] [] + >> Cases_on `SUC n' DIV 10` + >- (METIS_TAC [SUC_NOT,MULT_EQ_0, ADD_0, ADD_COMM, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ SRW_TAC [] [rec_toString_def] ++ MATCH_MP_TAC append_neq_lem - ++ RW_TAC list_ss []) - ++ Cases_on `rec_toString (SUC n' DIV 10)` - >> (SRW_TAC [] [CHAR_EQ_THM] - ++ `ORD (CHR (48 + SUC n' MOD 10)) = 48 + SUC n' MOD 10` + >> SRW_TAC [] [rec_toString_def] >> MATCH_MP_TAC append_neq_lem + >> RW_TAC list_ss []) + >> Cases_on `rec_toString (SUC n' DIV 10)` + >- (SRW_TAC [] [CHAR_EQ_THM] + >> `ORD (CHR (48 + SUC n' MOD 10)) = 48 + SUC n' MOD 10` by (RW_TAC std_ss [GSYM ORD_CHR] - ++ MATCH_MP_TAC LESS_LESS_EQ_TRANS ++ Q.EXISTS_TAC `48 + 10:num` - ++ RW_TAC arith_ss [LT_ADD_LCANCEL, + >> MATCH_MP_TAC LESS_LESS_EQ_TRANS >> Q.EXISTS_TAC `48 + 10:num` + >> RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ ONCE_REWRITE_TAC [ADD_COMM] - ++ RW_TAC arith_ss []) - ++ SPOSE_NOT_THEN STRIP_ASSUME_TAC - ++ `LENGTH (h::t ++ [CHR (SUC n' MOD 10 + 48)]) = LENGTH [#"0"]` + >> ONCE_REWRITE_TAC [ADD_COMM] + >> RW_TAC arith_ss []) + >> SPOSE_NOT_THEN STRIP_ASSUME_TAC + >> `LENGTH (h::t ++ [CHR (SUC n' MOD 10 + 48)]) = LENGTH [#"0"]` by ASM_SIMP_TAC bool_ss [] - ++ Q.PAT_ASSUM `h::t ++ [CHR (SUC n' MOD 10 + 48)] = [#"0"]` (K ALL_TAC) - ++ FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) - ++ EQ_TAC - >> (RW_TAC std_ss [toString_def, rec_toString_def, IMPLODE_11] - ++ ONCE_REWRITE_TAC [DECIDE ``(n' = n) = (SUC n' = SUC n)``] - ++ `(SUC n' = SUC n) = + >> Q.PAT_ASSUM `h::t ++ [CHR (SUC n' MOD 10 + 48)] = [#"0"]` (K ALL_TAC) + >> FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) + >> EQ_TAC + >- (RW_TAC std_ss [toString_def, rec_toString_def, IMPLODE_11] + >> ONCE_REWRITE_TAC [DECIDE ``(n' = n) = (SUC n' = SUC n)``] + >> `(SUC n' = SUC n) = ((SUC n') DIV 10 * 10 + (SUC n') MOD 10 = (SUC n) DIV 10 * 10 + (SUC n) MOD 10)` by METIS_TAC [MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)] - ++ POP_ORW - ++ FULL_SIMP_TAC std_ss [append_sing_eq_lem] - ++ (MP_TAC o Q.SPECL [`48 + SUC n' MOD 10`, `48 + SUC n MOD 10`]) CHR_11 - ++ `!n. 48 + SUC n MOD 10 < 256` - by (STRIP_TAC ++ MATCH_MP_TAC LESS_LESS_EQ_TRANS - ++ Q.EXISTS_TAC `48 + 10:num` - ++ RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP + >> POP_ORW + >> FULL_SIMP_TAC std_ss [append_sing_eq_lem] + >> (MP_TAC o Q.SPECL [`48 + SUC n' MOD 10`, `48 + SUC n MOD 10`]) CHR_11 + >> `!n. 48 + SUC n MOD 10 < 256` + by (STRIP_TAC >> MATCH_MP_TAC LESS_LESS_EQ_TRANS + >> Q.EXISTS_TAC `48 + 10:num` + >> RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ RW_TAC arith_ss [EQ_MULT_LCANCEL] - ++ Cases_on `SUC n' DIV 10` - >> (Cases_on `SUC n DIV 10` ++ RW_TAC arith_ss [] - ++ FULL_SIMP_TAC std_ss [rec_toString_def] - ++ `LENGTH ([]:char list) = LENGTH (rec_toString (SUC n'' DIV 10) + >> RW_TAC arith_ss [EQ_MULT_LCANCEL] + >> Cases_on `SUC n' DIV 10` + >- (Cases_on `SUC n DIV 10` >> RW_TAC arith_ss [] + >> FULL_SIMP_TAC std_ss [rec_toString_def] + >> `LENGTH ([]:char list) = LENGTH (rec_toString (SUC n'' DIV 10) ++ [CHR (48 + SUC n'' MOD 10)])` by METIS_TAC [] - ++ Q.PAT_ASSUM `[] = rec_toString (SUC n'' DIV 10) + >> Q.PAT_ASSUM `[] = rec_toString (SUC n'' DIV 10) ++ [CHR (48 + SUC n'' MOD 10)]` (K ALL_TAC) - ++ FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) - ++ Cases_on `SUC n DIV 10` - >> (RW_TAC arith_ss [] - ++ FULL_SIMP_TAC std_ss [rec_toString_def] - ++ `LENGTH ([]:char list) = + >> FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) + >> Cases_on `SUC n DIV 10` + >- (RW_TAC arith_ss [] + >> FULL_SIMP_TAC std_ss [rec_toString_def] + >> `LENGTH ([]:char list) = LENGTH (rec_toString (SUC n'' DIV 10) ++ [CHR (48 + SUC n'' MOD 10)])` by METIS_TAC [] - ++ Q.PAT_ASSUM `rec_toString (SUC n'' DIV 10) + >> Q.PAT_ASSUM `rec_toString (SUC n'' DIV 10) ++ [CHR (48 + SUC n'' MOD 10)] = []` (K ALL_TAC) - ++ FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) - ++ `SUC n' DIV 10 < SUC n'` - by (MATCH_MP_TAC DIV_LESS ++ RW_TAC arith_ss []) - ++ Suff `toString (SUC n'') = toString (SUC n''')` >> METIS_TAC [] - ++ SIMP_TAC arith_ss [toString_def, IMPLODE_11] - ++ ASM_REWRITE_TAC []) - ++ RW_TAC std_ss []); + >> FULL_SIMP_TAC arith_ss [LENGTH_APPEND, LENGTH, GSYM LENGTH_NIL]) + >> `SUC n' DIV 10 < SUC n'` + by (MATCH_MP_TAC DIV_LESS >> RW_TAC arith_ss []) + >> Suff `toString (SUC n'') = toString (SUC n''')` >- METIS_TAC [] + >> SIMP_TAC arith_ss [toString_def, IMPLODE_11] + >> ASM_REWRITE_TAC []) + >> RW_TAC std_ss []); val STRCAT_toString_inj = store_thm ("STRCAT_toString_inj", @@ -258,48 +251,48 @@ val toString_toNum_cancel = store_thm ("toString_toNum_cancel", ``!n. toNum(toString n) = n``, completeInduct_on `n` - ++ Cases_on `n` - >> (SRW_TAC [] [toString_def, toNum_def, rec_toNum_def]) - ++ SRW_TAC [] [toString_def, rec_toString_def, toNum_def, rec_toNum_def, + >> Cases_on `n` + >- (SRW_TAC [] [toString_def, toNum_def, rec_toNum_def]) + >> SRW_TAC [] [toString_def, rec_toString_def, toNum_def, rec_toNum_def, GSYM SNOC_APPEND, REVERSE_SNOC] - ++ (MP_TAC o Q.SPECL [`48 + SUC n' MOD 10`]) ORD_CHR - ++ `!n. 48 + SUC n MOD 10 < 256` - by (STRIP_TAC ++ MATCH_MP_TAC LESS_LESS_EQ_TRANS ++ Q.EXISTS_TAC `48 + 10:num` - ++ RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP + >> (MP_TAC o Q.SPECL [`48 + SUC n' MOD 10`]) ORD_CHR + >> `!n. 48 + SUC n MOD 10 < 256` + by (STRIP_TAC >> MATCH_MP_TAC LESS_LESS_EQ_TRANS >> Q.EXISTS_TAC `48 + 10:num` + >> RW_TAC arith_ss [LT_ADD_LCANCEL, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)]) - ++ RW_TAC arith_ss [] - ++ Suff `SUC n' MOD 10 + rec_toNum (REVERSE (rec_toString (SUC n' DIV 10))) 1 = + >> RW_TAC arith_ss [] + >> Suff `SUC n' MOD 10 + rec_toNum (REVERSE (rec_toString (SUC n' DIV 10))) 1 = SUC n' MOD 10 + (SUC n' DIV 10) * 10` - >> METIS_TAC [ADD_COMM, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)] - ++ RW_TAC arith_ss [EQ_ADD_RCANCEL] - ++ `SUC n' DIV 10 < SUC n'` - by (MATCH_MP_TAC DIV_LESS ++ RW_TAC arith_ss []) - ++ Suff `rec_toNum (REVERSE (rec_toString (SUC n' DIV 10))) 1 = + >- METIS_TAC [ADD_COMM, MATCH_MP DIVISION (DECIDE ``0:num < 10:num``)] + >> RW_TAC arith_ss [EQ_ADD_RCANCEL] + >> `SUC n' DIV 10 < SUC n'` + by (MATCH_MP_TAC DIV_LESS >> RW_TAC arith_ss []) + >> Suff `rec_toNum (REVERSE (rec_toString (SUC n' DIV 10))) 1 = 10 * (toNum (toString (SUC n' DIV 10)))` - >> METIS_TAC [] - ++ POP_ASSUM (K ALL_TAC) - ++ POP_ASSUM (K ALL_TAC) - ++ POP_ASSUM (K ALL_TAC) - ++ Cases_on `SUC n' DIV 10` - >> SRW_TAC [] [toString_def, toNum_def, rec_toNum_def, rec_toString_def] - ++ Q.ABBREV_TAC `Q = rec_toNum (REVERSE (rec_toString (SUC n))) 1` - ++ SRW_TAC [] [toString_def, toNum_def] - ++ Q.UNABBREV_TAC `Q` - ++ POP_ASSUM (K ALL_TAC) - ++ Suff `!n m. rec_toNum (REVERSE (rec_toString (SUC n))) (SUC m) = + >- METIS_TAC [] + >> POP_ASSUM (K ALL_TAC) + >> POP_ASSUM (K ALL_TAC) + >> POP_ASSUM (K ALL_TAC) + >> Cases_on `SUC n' DIV 10` + >- SRW_TAC [] [toString_def, toNum_def, rec_toNum_def, rec_toString_def] + >> Q.ABBREV_TAC `Q = rec_toNum (REVERSE (rec_toString (SUC n))) 1` + >> SRW_TAC [] [toString_def, toNum_def] + >> Q.UNABBREV_TAC `Q` + >> POP_ASSUM (K ALL_TAC) + >> Suff `!n m. rec_toNum (REVERSE (rec_toString (SUC n))) (SUC m) = 10 * rec_toNum (REVERSE (rec_toString (SUC n))) m` - >> METIS_TAC [SUC_0] - ++ completeInduct_on `SUC n` - ++ SRW_TAC [] [toString_def, rec_toString_def, toNum_def, rec_toNum_def, + >- METIS_TAC [SUC_0] + >> completeInduct_on `SUC n` + >> SRW_TAC [] [toString_def, rec_toString_def, toNum_def, rec_toNum_def, GSYM SNOC_APPEND, REVERSE_SNOC, LEFT_ADD_DISTRIB] - ++ ONCE_REWRITE_TAC [EXP] - ++ RW_TAC arith_ss [EQ_ADD_LCANCEL] - ++ `SUC n DIV 10 < SUC n` - by (MATCH_MP_TAC DIV_LESS ++ RW_TAC arith_ss []) - ++ Cases_on `SUC n DIV 10` - >> SRW_TAC [] [toString_def, toNum_def, rec_toNum_def, rec_toString_def] - ++ Q.PAT_ASSUM `!m. m < SUC n ==> !n. b` (MP_TAC o Q.SPECL [`SUC n''`]) - ++ RW_TAC std_ss []); + >> ONCE_REWRITE_TAC [EXP] + >> RW_TAC arith_ss [EQ_ADD_LCANCEL] + >> `SUC n DIV 10 < SUC n` + by (MATCH_MP_TAC DIV_LESS >> RW_TAC arith_ss []) + >> Cases_on `SUC n DIV 10` + >- SRW_TAC [] [toString_def, toNum_def, rec_toNum_def, rec_toString_def] + >> Q.PAT_ASSUM `!m. m < SUC n ==> !n. b` (MP_TAC o Q.SPECL [`SUC n''`]) + >> RW_TAC std_ss []); (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) diff --git a/examples/fermat/count/files.txt b/examples/fermat/count/files.txt deleted file mode 100644 index ae88eddd6a..0000000000 --- a/examples/fermat/count/files.txt +++ /dev/null @@ -1,36 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Count Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: November, 2020 *) -(* ------------------------------------------------------------------------- *) - -0 helperCount -- general enhancement about maps between sets. -* helperFunction -* binomial -* Euler -* necklace - -1 combinatorics -- combinations, permutations, and arrangements. -* 0 helperCount -* helperFunction -* binomial - -2 mapCount -- counting of maps between finite sets. -* 1 combinatorics -* 0 helperCount -* necklace - -3 symmetry -- symmetry group, symmetry field, and automorphisms. -* 2 mapCount -* 1 combinatorics -* 0 helperCount -* groupMap -* group - -4 permutation -- permutation group of lists. -* 3 symmetry -* 2 mapCount -* 1 combinatorics -* 0 helperCount -* group diff --git a/examples/fermat/little/files.txt b/examples/fermat/little/files.txt deleted file mode 100644 index f63c35adf0..0000000000 --- a/examples/fermat/little/files.txt +++ /dev/null @@ -1,68 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Fermat Little Theorem Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: October, 2020 *) -(* ------------------------------------------------------------------------- *) - -0 necklace -- necklaces, monocoloured and multicoloured -* helperNum -* helperSet -* helperFunction - -0 cycle -- cycles on lists. -* helperNum -* helperSet -* helperFunction - -1 pattern -- cycle period and order. -0 cycle - -0 FLTbinomial -- number-theoretic proof using binomial expansion of prime exponent. -* binomial -* divides - -0 FLTnumber -- using modular arithmetic, without explicit group. -* Euler -* helperFunction -* helperSet - -0 FLTgroup -- group-theoretic proof using multiplicative group of modulo prime. -* group -* groupOrder -* groupInstances - -0 FLTeuler -- group-theoretic proof using multiplicative group of coprimes. -* Euler -* group -* groupOrder -* groupInstances - -2 FLTnecklace -- using necklace cycles and patterns. -1 pattern -0 cycle -0 necklace - -3 FLTaction -- using group action on necklaces. -2 FLTnecklace -0 cycle -0 necklace -* group -* groupAction - -4 FLTfixedpoint -- using fixed points of group action on necklaces. -3 FLTaction -2 FLTnecklace -0 necklace -0 cycle -* group -* groupAction -* groupInstances - -4 FLTpetersen -- the original Petersen's proof, line-by-line. -3 FLTaction -2 FLTnecklace -0 necklace -* group -* groupAction -* groupInstances diff --git a/examples/fermat/twosq/README.md b/examples/fermat/twosq/README.md index 1537ee7ab8..0d3872114a 100644 --- a/examples/fermat/twosq/README.md +++ b/examples/fermat/twosq/README.md @@ -1,7 +1,11 @@ # Fermat's Two Squares Theorem, with an algorithm. -These theories support the formalisation of an algorithm for Fermat's Two Squares Theorem. This is the source code for the paper `Windmills of the Minds: an Algorithm for Fermat's Two Squares Theorem`, with [DOI: 10.6092/issn.1972-5787/3728](https://doi.org/10.1145/3497775.3503673) or [pre-print](https://arxiv.org/abs/2112.02556). +These theories support the formalisation of an algorithm for Fermat's Two Squares Theorem. +This library contain the source code for these papers: + +* `Windmills of the Minds: an Algorithm for Fermat's Two Squares Theorem`, with [DOI: 10.6092/issn.1972-5787/3728](https://doi.org/10.1145/3497775.3503673) or [pre-print](https://arxiv.org/abs/2112.02556). +* `Windmills of the Minds: a Hopping Algorithm for Fermat's Two Squares Theorem`, with [DOI: 10.1007/s10817-024-09708-3](https://doi.org/10.1007/s10817-024-09708-3) or [open access](https://rdcu.be/dXE1Y). ## Theory * __helperTwosq__, helper theorems for arithmetic, sets, maps and lists. diff --git a/examples/fermat/twosq/files.txt b/examples/fermat/twosq/files.txt deleted file mode 100644 index afc92cc2d5..0000000000 --- a/examples/fermat/twosq/files.txt +++ /dev/null @@ -1,72 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Fermat Two Squares Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: July, 2020 *) -(* ------------------------------------------------------------------------- *) - -0 helperTwosq -- helper for arithmetic, algebra, sets, maps and lists. -* helperNum -* helperSet -* helperFunction - -1 involute -- basic properties of involution. -0 helperTwosq - -2 involuteFix -- pairs and fixes of involution. -1 involute -0 helperTwosq - -3 windmill -- definitions of windmill, mill, flip, and zagier. -2 involuteFix -1 involute -0 helperTwosq - -2 involuteAction -- involution and group action. -1 involute -0 helperTwosq - -0 iteration -- iteration period of FUNPOW. - -1 iterateCompute -- iteration period computation, recursion and while-loop. -0 listRange -0 iteration -0 helperTwosq - -3 iterateCompose -- iteration of involution composition. -1 iterateCompute -2 involuteFix -1 involute -0 iteration -0 helperTwosq - -4 twoSquares -- existence, uniqueness and algorithm (clean version for paper). -2 involuteFix -1 involute -3 windmill -3 iterateCompose -1 iterateCompute -0 iteration -0 helperTwosq -* groupAction - -4 quarity -- bound and listing of mills and search for sum of two squares. -3 windmill -0 helperTwosq - -3 corner -- an adaptation of a partition proof using corner pieces. -2 quarity -2 involuteFix -1 involute -0 helperTwosq - -5 hopping -- path and nodes justifying the hopping algorithm for two squares. -3 iterateCompose -1 iterateCompute -0 iteration -2 involuteFix -1 involute -3 windmill -4 quarity -4 twoSquares -0 helperTwosq diff --git a/examples/formal-languages/FormalLangScript.sml b/examples/formal-languages/FormalLangScript.sml index 5bf82fef44..71461c673f 100644 --- a/examples/formal-languages/FormalLangScript.sml +++ b/examples/formal-languages/FormalLangScript.sml @@ -281,6 +281,12 @@ Proof RW_TAC basic_ss [IN_KSTAR] >> METIS_TAC [DOTn_def,IN_INSERT] QED +Triviality KSTAR_NONEMPTY: + KSTAR A ≠ ∅ +Proof + rw[EXTENSION] >> metis_tac[EPSILON_IN_KSTAR] +QED + Theorem KSTAR_SING: x ∈ KSTAR {x} Proof diff --git a/examples/formal-languages/regular/Holmakefile b/examples/formal-languages/regular/Holmakefile index 5d542e3497..171d71f841 100644 --- a/examples/formal-languages/regular/Holmakefile +++ b/examples/formal-languages/regular/Holmakefile @@ -1,6 +1,7 @@ INCLUDES = $(HOLDIR)/examples/balanced_bst \ $(HOLDIR)/examples/formal-languages \ - $(HOLDIR)/examples/formal-languages/context-free + $(HOLDIR)/examples/formal-languages/context-free \ + $(HOLDIR)/src/search all: $(DEFAULT_TARGETS) regexp2dfa diff --git a/examples/formal-languages/regular/regularScript.sml b/examples/formal-languages/regular/regularScript.sml index c86411ea5e..1150c231f3 100644 --- a/examples/formal-languages/regular/regularScript.sml +++ b/examples/formal-languages/regular/regularScript.sml @@ -1,7 +1,7 @@ (*===========================================================================*) (* Basic automata theory: nfas, dfas, their equivalence via the subset *) -(* construction, closure constructions, etc. The approach taken is *) -(* set-oriented, rather than computational. *) +(* construction, closure constructions, Myhill-Nerode, etc. The approach *) +(* taken is set-oriented, rather than computational. *) (*===========================================================================*) open HolKernel Parse boolLib bossLib; @@ -11,6 +11,8 @@ open combinTheory pairTheory listTheory rich_listTheory prim_recTheory arithmeticTheory FormalLangTheory; +open dirGraphTheory dftTheory; + infix byA; val op byA = BasicProvers.byA; @@ -66,6 +68,12 @@ pred_setTheory.X_LE_MAX_SET (THEOREM) (* Local lemmas, possibly of wider use. Start with sets *) (*---------------------------------------------------------------------------*) +Triviality forall_emptyset: + (∀x. x ∉ s) ⇔ s = ∅ +Proof + rw[EXTENSION] +QED + Theorem ELT_SUBSET: a ∈ s ⇔ {a} ⊆ s Proof @@ -108,6 +116,15 @@ Proof rw[] >> spose_not_then assume_tac >> drule_all X_LE_MAX_SET >> decide_tac QED +Triviality finite_image_range: + FINITE t ⇒ (∀x. x ∈ s ⇒ f x ∈ t) ⇒ FINITE(IMAGE f s) +Proof + rw [] >> + ‘IMAGE f s ⊆ t’ by + (rw[SUBSET_DEF] >> metis_tac[]) >> + irule SUBSET_FINITE >> metis_tac[] +QED + (*---------------------------------------------------------------------------*) (* List destructors: HD, TL, and LAST *) (*---------------------------------------------------------------------------*) @@ -708,7 +725,7 @@ Theorem is_exec_append: ⇒ is_exec N (qs1 ++ qs1) (w1 ++ w2) Proof - cheat + QED *) @@ -1460,6 +1477,12 @@ Proof rw [IN_REGULAR,SUBSET_DEF] >> metis_tac [wf_nfa_def] QED +Theorem REGULAR_IS_FORMAL_LANG: + (L,A) ∈ REGULAR ⇒ IS_FORMAL_LANG(L,A) +Proof + metis_tac [IS_FORMAL_LANG_def,REGULAR_BOUNDED,REGULAR_SIGMA_FINITE] +QED + Theorem EMPTYSET_IN_REGULAR: FINITE A ⇒ ({},A) ∈ REGULAR Proof @@ -2229,31 +2252,27 @@ Proof QED (*---------------------------------------------------------------------------*) -(* Closure under Kleene *) +(* Closure under Kleene Star. From *) +(* *) +(* wlist = [w1 ; ... ; wn] *) +(* *) +(* obtain *) +(* *) +(* qslist = [qs1 ; ... ; qsn] *) +(* *) +(* where *) +(* *) +(* is_accepting_exec qsi wi holds, for i < LENGTH wlist *) +(* *) +(* To then build the execution for the nfa_plus machine, we basically want *) +(* to build the full nfa_plus state list as "FLAT qslist". But, as for the *) +(* nfa_dot pasting lemma, we have to arrange for the branching back from *) +(* accept states to skip over the initial state. Thus for every qsi, *) +(* i > 1, drop the HD: *) +(* *) +(* nfa_plus_qs = qs1 ++ TL qs2 ++ ... ++ TL qsn *) (*---------------------------------------------------------------------------*) -(* - From - - wlist = [w1 ; ... ; wn] - - obtain - - qslist = [qs1 ; ... ; qsn] - - where - - is_accepting_exec qsi wi holds, for i < LENGTH wlist - - To then build the execution for the nfa_plus machine, we basically want - to build the full nfa_plus state list as "FLAT qslist". But, as for the - nfa_dot pasting lemma, we have to arrange for the branching back from - accept states to skip over the initial state. Thus for every qsi, i > 1, - drop the HD: - - nfa_plus_qs = qs1 ++ TL qs2 ++ ... ++ TL qsn -*) - Triviality qslist_length: ∀wlist qslist. LIST_REL (is_accepting_exec M) qslist wlist @@ -2411,44 +2430,6 @@ Proof irule is_accepting_exec_nfa_plus_paste_list >> rw[] QED -(* -Triviality SPLITP_EXISTS: - EXISTS P list ⇒ ∃prefix h t. SPLITP P list = (prefix,h::t) -Proof - Induct_on ‘list’ >> rw[SPLITP] >> gvs[] -QED - -Theorem zip_expand_id: - ZIP (MAP FST plist, MAP SND plist) = plist -Proof - Induct_on ‘plist’ >> rw[] -QED - -Theorem pair_list_as_zip: - ∀plist:('a # 'b) list. ∃l1 l2. LENGTH l1 = LENGTH l2 ∧ plist = ZIP (l1,l2) -Proof - Induct_on ‘plist’ >> rw[] - >- (ntac 2 (qexists_tac ‘[]’) >> rw[]) - >- (namedCases_on ‘h’ ["a b"] >> - qexists_tac ‘a::l1’ >> qexists_tac ‘b::l2’ >> rw[]) -QED - -Triviality UNZIP_APPEND: - ∀L1 L2. UNZIP (L1 ++ L2) = (FST (UNZIP L1) ++ FST (UNZIP L2), - SND (UNZIP L1) ++ SND (UNZIP L2)) -Proof - Induct >> rw[] -QED - -Theorem ZIP_EQ_NIL_OR: - ZIP(l1,l2) = [] ⇔ l1 = [] ∨ l2 = [] -Proof - Induct_on ‘l1’ >> Induct_on ‘l2’ >> rw[ZIP_def] -QED - - -*) - Triviality nfa_plus_diff: is_dfa M ∧ a ∈ M.Sigma ∧ @@ -2867,7 +2848,7 @@ Proof >- (qexists_tac ‘{q'}’ >> rw[]) QED -(* NB: The statement of this "Theorem" is not true: nfa_eval relies on +(* NB: The following Theorem statement is not true: nfa_eval relies on the input word having all its symbols drawn from the alphabet. The N.delta q a transition is not defined when a is not in the alphabet. But, in HOL such an expression has a "value"---an unspecified value---which, @@ -3155,7 +3136,7 @@ QED (*===========================================================================*) (*---------------------------------------------------------------------------*) -(* Work with relations relativized to sets, typically A* *) +(* Work with relations relativized to A* *) (*---------------------------------------------------------------------------*) Definition nfa_eval_equiv_def: @@ -3172,7 +3153,7 @@ Definition lang_equiv_def: (∀z. z ∈ KSTAR{[a] | a ∈ A} ⇒ (x ++ z ∈ L ⇔ y ++ z ∈ L)) End -Theorem nfa_eval_equiv_refines_lang_equiv: +Theorem nfa_eval_equiv_refines_lang_equiv[local]: wf_nfa N ∧ nfa_eval_equiv N x y ⇒ lang_equiv (nfa_lang N,N.Sigma) x y @@ -3217,16 +3198,16 @@ Proof rw[nfa_eval_append] QED +(*---------------------------------------------------------------------------*) +(* Lemmas on partitions *) +(*---------------------------------------------------------------------------*) + Theorem partition_def_alt: partition R s = {equiv_class R s x | x | x ∈ s} Proof rw[EXTENSION,partition_def,EQ_IMP_THM,PULL_EXISTS] >> metis_tac[] QED -(*---------------------------------------------------------------------------*) -(* Not used here, but perhaps useful somewhere *) -(*---------------------------------------------------------------------------*) - Theorem partition_def_as_image: partition R s = IMAGE (λx y. y ∈ s ∧ R x y) s Proof @@ -3234,6 +3215,58 @@ Proof simp[IN_DEF,ETA_THM] QED +Triviality in_partition: + class ∈ partition R s ⇔ + ∃x. x ∈ s ∧ ∀y. y ∈ class ⇔ y ∈ s ∧ R x y +Proof + rw [partition_def,EQ_IMP_THM,EXTENSION] +QED + +Triviality in_partition_alt: + class ∈ partition R s ⇔ ∃x. x ∈ s ∧ class = equiv_class R s x +Proof + rw [partition_def,EQ_IMP_THM,EXTENSION] +QED + +Triviality partition_empty: + partition R s = ∅ ⇔ s = ∅ +Proof + rw [EQ_IMP_THM] + >- (gvs [EXTENSION,in_partition] >> + gen_tac >> disch_tac >> gvs [GSYM IMP_DISJ_THM] >> + first_x_assum drule >> simp[] >> + qexists_tac ‘equiv_class R s x’ >> simp[]) + >- simp[partition_def] +QED + +Triviality kstar_partition_inhab_word: + E equiv_on (KSTAR{[a] | a ∈ A}) ∧ + class ∈ partition E (KSTAR{[a] | a ∈ A}) ⇒ + ∃w. w ∈ class ∧ equiv_class E (KSTAR{[a] | a ∈ A}) w = class +Proof + rw[in_partition_alt] >> qexists_tac ‘x’ >> gvs[equiv_on_def] +QED + +Theorem partition_elts: + R equiv_on s ∧ + t1 ∈ partition R s ∧ + t2 ∈ partition R s ∧ + w ∈ t1 ∧ w ∈ t2 ⇒ t1 = t2 +Proof + rw [in_partition_alt] >> + gvs[EXTENSION,equiv_on_def] >> + rw [EQ_IMP_THM] >> metis_tac[] +QED + +Triviality partition_emptylang: + partition (lang_equiv({},A)) (KSTAR{[a] | a ∈ A}) = {KSTAR{[a] | a ∈ A}} +Proof + simp [EXTENSION, in_partition, lang_equiv_def] >> rw [EQ_IMP_THM] + >- metis_tac[] + >- metis_tac[] + >- (qexists_tac ‘ε’ >> rw[]) +QED + (*---------------------------------------------------------------------------*) (* The set of words that evaluate to the given state. M will be a DFA. *) (*---------------------------------------------------------------------------*) @@ -3244,9 +3277,9 @@ Definition words_of_state_def: End Theorem in_words_of_state: - w ∈ words_of_state N q + w ∈ words_of_state M q ⇔ - EVERY (λa. a ∈ N.Sigma) w ∧ nfa_eval N N.initial w = {q} + EVERY (λa. a ∈ M.Sigma) w ∧ nfa_eval M M.initial w = {q} Proof rw[words_of_state_def] QED @@ -3264,25 +3297,17 @@ Proof QED (*---------------------------------------------------------------------------*) -(* An eval-partition is defined by words_of_state *) +(* Create an injective map from E-classes to machine states *) +(* *) +(* state_of_class_def: *) +(* |- ∀M. is_dfa M ⇒ *) +(* ∀class. *) +(* class ∈ partition (nfa_eval_equiv M) (KSTAR{[a] | a ∈ M.Sigma}) *) +(* ⇒ *) +(* state_of_class M class ∈ M.Q ∧ *) +(* class = words_of_state M (state_of_class M class) *) (*---------------------------------------------------------------------------*) -Theorem exists_state_class[local]: - is_dfa M ⇒ - ∀class. - class ∈ partition (nfa_eval_equiv M) (KSTAR{[a] | a ∈ M.Sigma}) - ⇒ ∃q. q ∈ M.Q ∧ class = words_of_state M q -Proof - rw [partition_def_alt] >> - rename1 ‘EVERY (λa. a ∈ M.Sigma) w’ >> - rw [words_of_state_def] >> - ‘∃q. nfa_eval M M.initial w = {q}’ by - metis_tac [dfa_eval_final_state] >> - rw [nfa_eval_equiv_def] >> - qexists_tac ‘q’ >> rw [EXTENSION] >> - metis_tac [dfa_eval_states_closed] -QED - Theorem state_of_class_witness[local]: ∃state_of_class. ∀M. is_dfa M ⇒ @@ -3292,23 +3317,17 @@ Theorem state_of_class_witness[local]: state_of_class M class ∈ M.Q ∧ class = words_of_state M (state_of_class M class) Proof - qexists_tac ‘λM class. @q. q ∈ M.Q ∧ class = words_of_state M q’ >> - rw[] >> SELECT_ELIM_TAC >> - metis_tac[exists_state_class] + rw[GSYM SKOLEM_THM,PUSH_EXISTS] >> + gvs [partition_def_alt] >> + rename1 ‘EVERY (λa. a ∈ M.Sigma) w’ >> + rw [words_of_state_def] >> + ‘∃q. nfa_eval M M.initial w = {q}’ by + metis_tac [dfa_eval_final_state] >> + rw [nfa_eval_equiv_def] >> + qexists_tac ‘q’ >> rw [EXTENSION] >> + metis_tac [dfa_eval_states_closed] QED -(*---------------------------------------------------------------------------*) -(* So can make a map from classes to states *) -(* *) -(* state_of_class_def: *) -(* |- ∀M. is_dfa M ⇒ *) -(* ∀class. *) -(* class ∈ partition (nfa_eval_equiv M) (KSTAR{[a] | a ∈ M.Sigma}) *) -(* ⇒ *) -(* state_of_class M class ∈ M.Q ∧ *) -(* class = words_of_state M (state_of_class M class) *) -(*---------------------------------------------------------------------------*) - val state_of_class_def = new_specification ("state_of_class_def", ["state_of_class"], state_of_class_witness); @@ -3345,8 +3364,8 @@ Proof QED (*---------------------------------------------------------------------------*) -(* The partition is finite because state_of_class is an injection from the *) -(* partition into the set of states. *) +(* The "nfa-eval"-partition is finite because state_of_class is an injection *) +(* from the partition into the (finite) set of states. *) (*---------------------------------------------------------------------------*) Theorem finite_nfa_eval_equiv_partition: @@ -3419,60 +3438,6 @@ Proof ) QED -Triviality forall_emptyset: - (∀x. x ∉ s) ⇔ s = ∅ -Proof - rw[EXTENSION] -QED - -Triviality in_partition: - class ∈ partition R s ⇔ - ∃x. x ∈ s ∧ ∀y. y ∈ class ⇔ y ∈ s ∧ R x y -Proof - rw [partition_def,EQ_IMP_THM,EXTENSION] -QED - -Triviality in_partition_alt: - class ∈ partition R s ⇔ ∃x. x ∈ s ∧ class = equiv_class R s x -Proof - rw [partition_def,EQ_IMP_THM,EXTENSION] -QED - -Triviality partition_empty: - partition R s = ∅ ⇔ s = ∅ -Proof - rw [EQ_IMP_THM] - >- (gvs [EXTENSION,in_partition] >> - gen_tac >> disch_tac >> gvs [GSYM IMP_DISJ_THM] >> - first_x_assum drule >> simp[] >> - qexists_tac ‘equiv_class R s x’ >> simp[]) - >- simp[partition_def] -QED - -Triviality KSTAR_NONEMPTY: - KSTAR A ≠ ∅ -Proof - rw[EXTENSION] >> metis_tac[EPSILON_IN_KSTAR] -QED - -Triviality kstar_partition_inhab_word: - E equiv_on (KSTAR{[a] | a ∈ A}) ∧ - class ∈ partition E (KSTAR{[a] | a ∈ A}) ⇒ - ∃w. w ∈ class ∧ equiv_class E (KSTAR{[a] | a ∈ A}) w = class -Proof - rw[in_partition_alt] >> qexists_tac ‘x’ >> - gvs[equiv_on_def] -QED - -Triviality partition_emptylang: - partition (lang_equiv({},A)) (KSTAR{[a] | a ∈ A}) = {KSTAR{[a] | a ∈ A}} -Proof - simp [EXTENSION, in_partition, lang_equiv_def] >> rw [EQ_IMP_THM] - >- metis_tac[] - >- metis_tac[] - >- (qexists_tac ‘ε’ >> rw[]) -QED - (*---------------------------------------------------------------------------*) (* Every right-invariant equivalence on A* is a refinement of language *) (* equivalence. *) @@ -3498,8 +3463,7 @@ Proof metis_tac [SUBSET_DEF] >> pop_keep_tac >> rw[in_partition] >- (‘EVERY (λa. a ∈ A) (x++z) ∧ E x' (x++z)’ by metis_tac[] >> - simp[] >> - qpat_x_assum ‘_ equiv_on _’ mp_tac >> + simp[] >> qpat_x_assum ‘_ equiv_on _’ mp_tac >> simp[equiv_on_def] >> disch_then (irule o last o CONJUNCTS) >> rw[] >> first_x_assum (irule_at Any) >> simp[]) @@ -3525,27 +3489,35 @@ Proof metis_tac [SUBSET_DEF] QED -Triviality finite_image_range: - FINITE t ⇒ (∀x. x ∈ s ⇒ f x ∈ t) ⇒ FINITE(IMAGE f s) -Proof - rw [] >> - ‘IMAGE f s ⊆ t’ by - (rw[SUBSET_DEF] >> metis_tac[]) >> - irule SUBSET_FINITE >> metis_tac[] -QED +(*---------------------------------------------------------------------------*) +(* Create an injective map from L-classes to (some) E-classes *) +(* *) +(* E_class_def: *) +(* |- ∀A E classes class'. *) +(* E equiv_on KSTAR{[a] | a ∈ A} ∧ *) +(* right_invar (KSTAR {[a] | a ∈ A}) E ∧ *) +(* classes ⊆ partition E (KSTAR {[a] | a ∈ A}) ∧ *) +(* class' ∈ partition *) +(* (lang_equiv (BIGUNION classes,A)) (KSTAR {[a] | a ∈ A}) *) +(* ==> *) +(* E_class A E classes class' ∈ partition E (KSTAR {[a] | a ∈ A}) ∧ *) +(* E_class A E classes class' ⊆ class' *) +(*---------------------------------------------------------------------------*) -Triviality from_above: - ∀A classes class'. +Triviality E_class_witness: + ∃E_part. + ∀A E classes class'. E equiv_on (KSTAR{[a] | a ∈ A}) ∧ right_invar (KSTAR{[a] | a ∈ A}) E ∧ classes ⊆ partition E (KSTAR{[a] | a ∈ A}) ∧ class' ∈ partition (lang_equiv(BIGUNION classes,A)) (KSTAR{[a] | a ∈ A}) ⇒ - ∃p. p ∈ partition E (KSTAR{[a] | a ∈ A}) ∧ p ⊆ class' + E_part A E classes class' ∈ partition E (KSTAR{[a] | a ∈ A}) ∧ + E_part A E classes class' ⊆ class' Proof - rw[] >> - ‘lang_equiv (BIGUNION classes,A) equiv_on KSTAR{[a] | a ∈ A}’ by - metis_tac[equiv_on_lang_equiv] >> + rw [GSYM SKOLEM_THM,PUSH_EXISTS] >> + ‘lang_equiv (BIGUNION classes,A) equiv_on KSTAR{[a] | a ∈ A}’ by + metis_tac[equiv_on_lang_equiv] >> drule_all kstar_partition_inhab_word >> strip_tac >> ‘w ∈ KSTAR{[a] | a ∈ A}’ by gvs [in_partition_alt] >> @@ -3554,40 +3526,8 @@ Proof gvs[] >> metis_tac[] QED -Triviality from_above_witness: - ∃E_part. - ∀A E classes class'. - E equiv_on (KSTAR{[a] | a ∈ A}) ∧ - right_invar (KSTAR{[a] | a ∈ A}) E ∧ - classes ⊆ partition E (KSTAR{[a] | a ∈ A}) ∧ - class' ∈ partition (lang_equiv(BIGUNION classes,A)) (KSTAR{[a] | a ∈ A}) - ⇒ - E_part A E classes class' ∈ partition E (KSTAR{[a] | a ∈ A}) ∧ - E_part A E classes class' ⊆ class' -Proof - qexists_tac ‘λA E classes class'. - @part. part ∈ partition E (KSTAR{[a] | a ∈ A}) ∧ - part ⊆ class'’ >> - rw[] >> SELECT_ELIM_TAC >> metis_tac[from_above] -QED - -(*---------------------------------------------------------------------------*) -(* So can make an injective map from L-classes to some E-classes *) -(* *) -(* E_class_def: *) -(* |- ∀A E classes class'. *) -(* E equiv_on KSTAR{[a] | a ∈ A} ∧ *) -(* right_invar (KSTAR {[a] | a ∈ A}) E ∧ *) -(* classes ⊆ partition E (KSTAR {[a] | a ∈ A}) ∧ *) -(* class' ∈ partition *) -(* (lang_equiv (BIGUNION classes,A)) (KSTAR {[a] | a ∈ A}) *) -(* ==> *) -(* E_class A E classes class' ∈ partition E (KSTAR {[a] | a ∈ A}) ∧ *) -(* E_class A E classes class' ⊆ class' *) -(*---------------------------------------------------------------------------*) - val E_class_def = - new_specification ("E_class_def", ["E_class"], from_above_witness); + new_specification ("E_class_def", ["E_class"], E_class_witness); Theorem E_class_inj: E equiv_on (KSTAR{[a] | a ∈ A}) ∧ @@ -3612,9 +3552,9 @@ QED (*---------------------------------------------------------------------------*) (* The injection is into a subset of the E-partition, which is finite. Thus *) -(* the source L-partition is finite. NB: I was confused for a while about *) -(* the role of the "classes" subset of the E-partition in the argument, but *) -(* it doesn't figure at this stage. (It does get used in the crucial *) +(* the source L-partition is finite. NB: one can become confused about the *) +(* role of the "classes" subset of the E-partition, but it doesn't figure at *) +(* this stage of the argument. (It is only needed in the crucial *) (* lang_equiv_refinement lemma.) *) (*---------------------------------------------------------------------------*) @@ -3628,12 +3568,258 @@ Theorem Myhill_Nerode_B: FINITE (partition (lang_equiv(L,A)) (KSTAR{[a] | a ∈ A})) Proof rw[] >> - irule (iffLR (FINITE_IMAGE_INJ_EQ - |> Q.ISPEC ‘f : (α list -> bool) -> (α list -> bool)’)) >> + irule (iffLR + (FINITE_IMAGE_INJ_EQ + |> Q.ISPEC ‘f : (α list -> bool) -> (α list -> bool)’)) >> qexists_tac ‘E_class A E classes’ >> rw[] >- (irule E_class_inj >> metis_tac[]) - >- (irule finite_image_range >> first_x_assum (irule_at Any) >> - rw[E_class_def]) + >- (irule finite_image_range >> + first_x_assum (irule_at Any) >> rw[E_class_def]) +QED + +val cong_tac = REFL_TAC ORELSE MK_COMB_TAC ORELSE ABS_TAC + +Triviality lang_equiv_abs: + lang_equiv (L,A) x y + ⇒ + lang_equiv (L,A) x = lang_equiv (L,A) y +Proof + strip_tac >> + ‘lang_equiv (L,A) equiv_on KSTAR{[a] | a ∈ A}’ by + metis_tac[equiv_on_lang_equiv] >> + ‘x ∈ KSTAR{[a] | a ∈ A} ∧ + y ∈ KSTAR{[a] | a ∈ A}’ by + metis_tac [lang_equiv_def] >> + rw[EXTENSION,IN_DEF,EQ_IMP_THM] >> + ‘x' ∈ KSTAR{[a] | a ∈ A}’ by + metis_tac [lang_equiv_def] >> + gvs [equiv_on_def] >> metis_tac[] +QED + +(*---------------------------------------------------------------------------*) +(* Construct a DFA from a finite lang_equiv partition of A*. States are, *) +(* again, not numbers in this construction, so work via a bijection to a *) +(* "count" set. *) +(*---------------------------------------------------------------------------*) + +Theorem Myhill_Nerode_C: + IS_FORMAL_LANG(L,A) ∧ + FINITE (partition (lang_equiv(L,A)) (KSTAR{[a] | a ∈ A})) + ⇒ + (L,A) ∈ REGULAR +Proof + rw[IN_REGULAR, IS_FORMAL_LANG_def] >> + qabbrev_tac ‘Qparts = partition (lang_equiv (L,A)) (KSTAR {[a] | a ∈ A})’ >> + ‘∃partOf k. BIJ partOf (count k) Qparts’ by metis_tac [FINITE_BIJ_COUNT] >> + imp_res_tac BIJ_LINV_INV >> + qabbrev_tac ‘numOf = LINV partOf (count k)’ >> + ‘INJ numOf Qparts (count k)’ by + (imp_res_tac BIJ_LINV_BIJ >> gvs [BIJ_DEF]) >> + qabbrev_tac ‘Lclass = equiv_class (lang_equiv (L,A)) (KSTAR {[a] | a ∈ A})’ >> + ‘lang_equiv (L,A) equiv_on KSTAR{[a] | a ∈ A}’ by + metis_tac[equiv_on_lang_equiv] >> + ‘∃start. start ∈ Qparts ∧ start = Lclass ε’ by + (qexists_tac ‘Lclass ε’ >> simp[] >> + qunabbrev_tac ‘Qparts’ >> + rw [in_partition_alt] >> + irule_at Any EQ_REFL >> simp[]) >> + ‘∃finals. finals ⊆ Qparts ∧ finals = IMAGE Lclass L’ by + (qexists_tac ‘IMAGE Lclass L’ >> rw[] >> + qunabbrev_tac ‘Lclass’ >> + qunabbrev_tac ‘Qparts’ >> + gvs [SUBSET_DEF] >> rw [in_partition_alt] >> metis_tac[]) >> + qexists_tac + ‘<| Sigma := A; + Q := IMAGE numOf Qparts; + initial := {numOf start}; + final := IMAGE numOf finals; + delta := (λnq a. {numOf (Lclass ((@w. w ∈ partOf nq) ++ [a]))}) + |>’ >> simp[] >> conj_asm1_tac + >- (rw [wf_nfa_def] >> rw[] >> irule_at Any EQ_REFL >> + SELECT_ELIM_TAC >> rename1 ‘state ∈ Qparts’ >> + qunabbrev_tac ‘Qparts’ >> + drule_all kstar_partition_inhab_word >> rw[] + >- metis_tac[] + >- (rw[in_partition] >> + qunabbrev_tac ‘Lclass’ >> gvs[] >> + qexists_tac ‘w ++ [a]’ >> rw[] >> + (cong_tac >> cong_tac >> TRY REFL_TAC) >> + ‘lang_equiv (L,A) (w ⧺ [a]) (x' ⧺ [a])’ by + (mp_tac (right_invar_lang_equiv |> INST_TYPE[alpha |-> “:num”]) >> + rw [right_invar_def] >> first_x_assum irule >> simp[]) >> + metis_tac [lang_equiv_abs])) + >> + qabbrev_tac + ‘M = <|Q := IMAGE numOf Qparts; Sigma := A; + delta := (λnq a. {numOf (Lclass ((@w. w ∈ partOf nq) ⧺ [a]))}); + initial := {numOf (Lclass ε)}; final := IMAGE numOf (IMAGE Lclass L) |>’ >> + (* Following lemma could be stated and proved separately, but we would have + to pull a lot of context out to state it. *) + ‘∀x y. x ∈ KSTAR {[a] | a ∈ A} ∧ y ∈ KSTAR {[a] | a ∈ A} + ⇒ nfa_eval M {numOf(Lclass x)} y = {numOf (Lclass (x++y))}’ by + (gen_tac >> recInduct SNOC_INDUCT >> rw[] + >- rw[nfa_eval_eqns] >> + gvs [EVERY_SNOC] >> rw [SNOC_APPEND] >> + drule nfa_eval_append >> rename1 ‘a ∈ A’ >> + disch_then (mp_tac o Q.SPEC ‘[a]’ o Q.SPEC ‘l’) >> + ‘{numOf (Lclass x)} ⊆ M.Q’ by + (‘M.Q = IMAGE numOf Qparts’ by gvs[Abbr‘M’] >> + simp[SUBSET_DEF] >> irule_at Any EQ_REFL >> + qunabbrev_tac ‘Qparts’ >> rw [in_partition_alt] >> + metis_tac[]) >> + ‘M.Sigma = A’ by gvs[Abbr‘M’] >> pop_subst_tac >> + simp[] >> disch_then kall_tac >> rw[nfa_eval_eqns] >> + ‘M.delta = λnq a. {numOf (Lclass ((@w. w ∈ partOf nq) ⧺ [a]))}’ by + gvs [Abbr‘M’] >> pop_subst_tac >> + simp[Once EXTENSION] >> gen_tac >> + irule (METIS_PROVE[] “A = B ⇒ (x = A ⇔ x = B)”) >> + AP_TERM_TAC >> + ‘Lclass (x ⧺ l) ∈ Qparts’ by + (qunabbrev_tac ‘Qparts’ >> + rw [in_partition_alt] >> + irule_at Any EQ_REFL >> simp[]) >> + SELECT_ELIM_TAC >> rw[] + >- (qexists_tac ‘x ++ l’ >> rw [Abbr ‘Lclass’] >> + ‘(x ++ l) ∈ KSTAR {[a] | a ∈ A}’ by simp [] >> + metis_tac [EVERY_APPEND,equiv_on_def]) + >- (pop_keep_tac >> qunabbrev_tac ‘Lclass’ >> + rw[EXTENSION] >> (ntac 2 cong_tac >> TRY REFL_TAC) >> + irule lang_equiv_abs >> + irule (iffLR right_invar_def) >> conj_tac + >- gvs [equiv_on_def] + >- (irule_at Any right_invar_lang_equiv >> rw[])) + ) + >> + pop_assum (mp_tac o Q.SPEC ‘ε’) >> rw[] >> + simp [EXTENSION,in_nfa_lang_nfa_eval_alt] >> + ‘M.Sigma = A ∧ M.initial = {numOf (Lclass ε)}’ by + gvs[Abbr‘M’] >> ntac 2 pop_subst_tac >> + rw [EQ_IMP_THM,PULL_EXISTS] + >- (‘EVERY (λa. a ∈ A) x’ by + gvs[SUBSET_DEF] >> + qexists_tac ‘numOf (Lclass x)’ >> rw[] >> + qunabbrev_tac ‘M’ >> rw[]) + >- (ntac 2 pop_keep_tac >> rw[Abbr‘M’] >> + rename1 ‘numOf (Lclass x) = numOf (Lclass y)’ >> + ‘Lclass x ∈ Qparts ∧ Lclass y ∈ Qparts’ by + (qunabbrev_tac ‘Qparts’ >> gvs [SUBSET_DEF] >> + rw[in_partition_alt] >> metis_tac[]) >> + ‘Lclass x = Lclass y’ by + metis_tac [INJ_DEF] >> + ‘EVERY (λa. a ∈ A) y’ by + (gvs[SUBSET_DEF] >> metis_tac[]) >> + drule (iffLR equiv_class_eq) >> simp[] >> + disch_then drule_all >> + simp[lang_equiv_def] >> + metis_tac [EVERY_DEF,APPEND_NIL]) +QED + +Theorem Myhill_Nerode: + ((L,A) ∈ REGULAR + ⇒ FINITE (partition (lang_equiv (L,A)) (KSTAR {[a] | a ∈ A}))) + ∧ + (IS_FORMAL_LANG (L,A) ∧ + FINITE (partition (lang_equiv (L,A)) (KSTAR {[a] | a ∈ A})) + ⇒ + (L,A) ∈ REGULAR) +Proof + metis_tac [Myhill_Nerode_A,Myhill_Nerode_B,Myhill_Nerode_C] +QED + +(*===========================================================================*) +(* DFA state minimization. This is accomplished by, first, removing all *) +(* non-accessible states, then coalescing all non-distinguishable states. *) +(* Both of these operations preserve the language originally recognized. *) +(*===========================================================================*) + +Definition accessible_state_def: + accessible_state M q ⇔ + ∃w. w ∈ KSTAR {[a] | a ∈ M.Sigma} ∧ + nfa_eval M M.initial w = {q} +End + +Definition accessible_nfa_def: + accessible_nfa M ⇔ ∀q. q ∈ M.Q ⇒ accessible_state M q +End + +(*---------------------------------------------------------------------------*) +(* States q1 and q2 are distinguishable if there is at least one string w *) +(* s.t. executing M on w from one of q1 or q2 ends in an accepting state but *) +(* executing from the other ends in a non-accepting state. *) +(*---------------------------------------------------------------------------*) + +Definition distinguishable_states_def: + distinguishable_states M q1 q2 ⇔ + {q1; q2} ⊆ M.Q ∧ + ∃w. w ∈ KSTAR {[a] | a ∈ M.Sigma} ∧ + (nfa_eval M {q1} w ⊆ M.final + ⇔ + ~(nfa_eval M {q2} w ⊆ M.final)) +End + +Definition distinguishable_nfa_def: + distinguishable_nfa M ⇔ + ∀q1 q2. {q1; q2} ⊆ M.Q ∧ q1 ≠ q2 ⇒ distinguishable_states M q1 q2 +End + +(*---------------------------------------------------------------------------*) +(* Accessibility is computed by an instance of depth-first traversal. *) +(*---------------------------------------------------------------------------*) + +Definition kidlist_def: + kidlist N q = + if q ∈ N.Q then + SET_TO_LIST(BIGUNION{N.delta q a | a | a ∈ N.Sigma}) + else [] +End + +Theorem nfa_parents_finite: + wf_nfa N ⇒ FINITE (Parents (kidlist N)) +Proof + strip_tac >> irule SUBSET_FINITE >> + qexists_tac ‘N.Q’ >> + gvs [wf_nfa_def] >> + rw[Parents_def, kidlist_def,SUBSET_DEF] >> + pop_keep_tac >> IF_CASES_TAC >> rw[] QED +(*---------------------------------------------------------------------------*) +(* Invocation: reachable_states N [] (SET_TO_LIST N.initial) [] *) +(*---------------------------------------------------------------------------*) + +Definition reachable_states_def: + reachable_states N = DFT (kidlist N) (list$CONS) +End + +Theorem reachable_states_eqns: + wf_nfa N + ⇒ + reachable_states N seen [] acc = acc ∧ + reachable_states N seen (h::t) acc = + if MEM h seen + then reachable_states N seen t acc + else reachable_states N (h::seen) (kidlist N h ++ t) (h::acc) +Proof + PURE_REWRITE_TAC[reachable_states_def] >> strip_tac >> + imp_res_tac nfa_parents_finite >> + rw[DFT_def] +QED + +(*---------------------------------------------------------------------------*) +(* A state is reachable iff reachable_states finds it *) +(*---------------------------------------------------------------------------*) + +Theorem reachable_states_spec: + wf_nfa N ⇒ + ∀q. q ∈ REACH_LIST (kidlist N) (SET_TO_LIST N.initial) + ⇔ + q ∈ set(reachable_states N [] (SET_TO_LIST N.initial) []) +Proof + PURE_REWRITE_TAC[reachable_states_def] >> + strip_tac >> + imp_res_tac nfa_parents_finite >> + irule DFT_REACH_THM >> metis_tac[] +QED + + val _ = export_theory(); diff --git a/examples/l3-machine-code/common/temporal_stateScript.sml b/examples/l3-machine-code/common/temporal_stateScript.sml index 52a0d2f942..7f081ab101 100644 --- a/examples/l3-machine-code/common/temporal_stateScript.sml +++ b/examples/l3-machine-code/common/temporal_stateScript.sml @@ -3,9 +3,6 @@ open set_sepTheory progTheory temporalTheory stateTheory val _ = new_theory "temporal_state" -infix \\ -val op \\ = op THEN; - (* ------------------------------------------------------------------------ *) val TEMPORAL_NEXT_def = Define` diff --git a/examples/lambda/barendregt/boehmScript.sml b/examples/lambda/barendregt/boehmScript.sml index adaf2c963c..1221ec89d7 100644 --- a/examples/lambda/barendregt/boehmScript.sml +++ b/examples/lambda/barendregt/boehmScript.sml @@ -11,12 +11,12 @@ open HolKernel Parse boolLib bossLib; open optionTheory arithmeticTheory pred_setTheory listTheory rich_listTheory llistTheory relationTheory ltreeTheory pathTheory posetTheory hurdUtils pairTheory finite_mapTheory topologyTheory listRangeTheory combinTheory - tautLib listLib string_numTheory; + tautLib listLib string_numTheory numLib BasicProvers; (* local theories *) open basic_swapTheory nomsetTheory termTheory appFOLDLTheory NEWLib chap2Theory chap3Theory reductionEval head_reductionTheory head_reductionLib - standardisationTheory solvableTheory horeductionTheory takahashiS3Theory; + standardisationTheory solvableTheory; (* enable basic monad support *) open monadsyntax; @@ -49,29 +49,6 @@ fun PRINT_TAC pfx g = (print (pfx ^ "\n"); ALL_TAC g); Overload FV = “supp term_pmact” Overload VAR = “term$VAR” -(*---------------------------------------------------------------------------* - * ltreeTheory extras - *---------------------------------------------------------------------------*) - -(* ltree_subset A B <=> A results from B by "cutting off" some subtrees. Thus, - - 1) The paths of A is a subset of paths of B - 2) The node contents for all paths of A is identical to those of B, but the number - of successors at those nodes of B may be different (B may have more successors) - - NOTE: Simply defining ‘ltree_subset A B <=> subtrees A SUBSET subtrees B’ is wrong, - as A may appear as a non-root subtree of B, i.e. ‘A IN subtrees B’ but there's - no way to "cut off" B to get A, the resulting subtree in B always have some - more path prefixes. - *) -Definition ltree_subset_def : - ltree_subset A B <=> - (ltree_paths A) SUBSET (ltree_paths B) /\ - !p. p IN ltree_paths A ==> - ltree_node (THE (ltree_lookup A p)) = - ltree_node (THE (ltree_lookup B p)) -End - (*---------------------------------------------------------------------------* * Boehm Trees (and subterms) - name after Corrado_Böhm [2] UOK * *---------------------------------------------------------------------------*) @@ -444,13 +421,13 @@ Proof >> qunabbrev_tac ‘M'’ >> qabbrev_tac ‘Y = RANK r’ >> qabbrev_tac ‘Y' = RANK (SUC r)’ - (* #1 *) + (* transitivity no.1 *) >> Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M1’ >> CONJ_TAC >- (FIRST_X_ASSUM MATCH_MP_TAC >> rw []) - (* #2 *) + (* transitivity no.2 *) >> Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M0 UNION set vs’ >> CONJ_TAC >- simp [FV_LAMl] - (* #3 *) + (* transitivity no.3 *) >> Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M UNION set vs’ >> CONJ_TAC >- (Suff ‘FV M0 SUBSET FV M’ >- SET_TAC [] \\ @@ -611,6 +588,9 @@ Proof >> MATCH_MP_TAC EL_MEM >> art [] QED +(* NOTE: This lemma is suitable for doing induction. A better upper bound is + given by the next [subterm_headvar_lemma']. + *) Theorem subterm_headvar_lemma : !X M r M0 n vs M1. FINITE X /\ FV M SUBSET X UNION RANK r /\ @@ -657,36 +637,70 @@ Proof >> rw [RNEWS_SUBSET_RANK] QED -Theorem subterm_headvar_lemma' : - !X M r M0 n vs M1. +(* NOTE: If ‘vs’ is longer than ‘n’, then ‘hnf_headvar M1’ uses at most + variables from ‘set (TAKE n vs). This result generalizes above lemmas. + *) +Theorem subterm_headvar_lemma_alt : + !X M r M0 n n' vs M1. FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M /\ M0 = principle_hnf M /\ n = LAMl_size M0 /\ - vs = RNEWS r n X /\ + vs = RNEWS r n' X /\ n <= n' /\ M1 = principle_hnf (M0 @* MAP VAR vs) - ==> hnf_headvar M1 IN FV M UNION set vs + ==> hnf_headvar M1 IN FV M UNION set (TAKE n vs) Proof RW_TAC std_ss [] >> qabbrev_tac ‘M0 = principle_hnf M’ >> qabbrev_tac ‘n = LAMl_size M0’ - >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ + >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n' :num”)) ‘X’ >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ >> ‘DISJOINT (set vs) (FV M0)’ by PROVE_TAC [subterm_disjoint_lemma'] >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, “y :string”, “args :term list”)) ‘M1’ - >> ‘TAKE n vs = vs’ by rw [] - >> POP_ASSUM (rfs o wrap) + >> rfs [] + >> qabbrev_tac ‘xs = TAKE n vs’ + >> Q.PAT_X_ASSUM ‘M1 = _’ K_TAC + >> Q.PAT_X_ASSUM ‘M0 = _’ (ASSUME_TAC o SYM) >> simp [] + (* redefine M1 *) + >> qunabbrev_tac ‘M1’ + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ + (* applying principle_hnf_beta_reduce_ext *) + >> Know ‘M1 = VAR y @* args @* DROP n (MAP VAR vs)’ + >- (qunabbrev_tac ‘M1’ \\ + qabbrev_tac ‘t = VAR y @* args’ \\ + rw [GSYM MAP_DROP] \\ + Know ‘MAP VAR vs = MAP VAR xs ++ MAP VAR (DROP n vs)’ + >- (REWRITE_TAC [GSYM MAP_APPEND] >> AP_TERM_TAC \\ + rw [Abbr ‘xs’, TAKE_DROP]) >> Rewr' \\ + REWRITE_TAC [appstar_APPEND] \\ + qabbrev_tac ‘l = MAP VAR (DROP n vs)’ \\ + MATCH_MP_TAC principle_hnf_beta_reduce_ext \\ + rw [Abbr ‘t’, hnf_appstar]) + >> DISCH_TAC >> ‘hnf M1’ by rw [hnf_appstar] - >> Q.PAT_X_ASSUM ‘M0 = _’ (ASSUME_TAC o SYM) >> Q.PAT_X_ASSUM ‘M1 = _’ (ASSUME_TAC o SYM) - >> Know ‘FV (principle_hnf (M0 @* MAP VAR vs)) SUBSET FV (M0 @* MAP VAR vs)’ + >> Know ‘hnf_headvar M1 = y’ + >- (POP_ASSUM (REWRITE_TAC o wrap o SYM) \\ + simp [GSYM appstar_APPEND]) + >> Rewr' + >> qabbrev_tac ‘M1' = principle_hnf (M0 @* MAP VAR xs)’ + >> Know ‘M1' = VAR y @* args’ + >- (qunabbrev_tac ‘M1'’ \\ + qabbrev_tac ‘t = VAR y @* args’ \\ + Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ + MATCH_MP_TAC principle_hnf_beta_reduce \\ + rw [Abbr ‘t’, hnf_appstar]) + >> DISCH_THEN (ASSUME_TAC o SYM) + >> Know ‘FV (principle_hnf (M0 @* MAP VAR xs)) SUBSET FV (M0 @* MAP VAR xs)’ >- (MATCH_MP_TAC principle_hnf_FV_SUBSET' \\ - ‘solvable M1’ by rw [solvable_iff_has_hnf, hnf_has_hnf] \\ - Suff ‘M0 @* MAP VAR vs == M1’ >- PROVE_TAC [lameq_solvable_cong] \\ - rw []) + ‘solvable M1'’ by rw [solvable_iff_has_hnf, hnf_has_hnf, hnf_appstar] \\ + Suff ‘M0 @* MAP VAR xs == M1'’ >- PROVE_TAC [lameq_solvable_cong] \\ + Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ + simp []) >> simp [] - >> POP_ASSUM (ONCE_REWRITE_TAC o wrap o SYM) + (* expand M1' *) + >> POP_ASSUM (REWRITE_TAC o wrap o SYM) >> simp [FV_appstar] >> STRIP_TAC >> Q.PAT_X_ASSUM ‘y IN _’ MP_TAC @@ -695,6 +709,26 @@ Proof >> MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art [] QED +Theorem subterm_headvar_lemma' : + !X M r M0 n vs M1. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + solvable M /\ + M0 = principle_hnf M /\ + n = LAMl_size M0 /\ + vs = RNEWS r n X /\ + M1 = principle_hnf (M0 @* MAP VAR vs) + ==> hnf_headvar M1 IN FV M UNION set vs +Proof + RW_TAC std_ss [] + >> qabbrev_tac ‘M0 = principle_hnf M’ + >> qabbrev_tac ‘n = LAMl_size M0’ + >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ + >> ‘vs = TAKE n vs’ by rw [] >> POP_ORW + >> MATCH_MP_TAC subterm_headvar_lemma_alt + >> qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’] >> simp [] +QED + (* Proposition 10.1.6 [1, p.219] *) Theorem lameq_BT_cong : !X M N r. FINITE X /\ FV M UNION FV N SUBSET X /\ M == N ==> @@ -951,51 +985,68 @@ Proof >> Q.PAT_X_ASSUM ‘Ms = Ms'’ (fs o wrap o SYM) QED -(* NOTE: When subterm X M p = NONE, either 1) M or its subterm is unsolvable, - or 2) p runs out of ltree_paths (BT X M). If we knew subterm X M p <> NONE - a priori, then p IN ltree_paths (BT X M) must hold. - *) -Theorem subterm_imp_ltree_paths : - !p X M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - subterm X M p r <> NONE ==> - p IN ltree_paths (BT' X M r) +Theorem BT_ltree_el_eq_none : + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r ==> + (ltree_el (BT' X M r) p = NONE <=> subterm X M p r = NONE) Proof - Induct_on ‘p’ >- rw [] - >> rpt GEN_TAC - >> STRIP_TAC - >> POP_ASSUM MP_TAC (* subterm X M (h::p) r <> NONE *) + Suff ‘!X. FINITE X ==> + !p M r. FV M SUBSET X UNION RANK r ==> + (ltree_el (BT' X M r) p = NONE <=> + subterm X M p r = NONE)’ + >- METIS_TAC [] + >> Q.X_GEN_TAC ‘X’ >> DISCH_TAC + >> Induct_on ‘p’ + >- rw [BT_ltree_el_NIL] + >> rpt STRIP_TAC >> reverse (Cases_on ‘solvable M’) - >- simp [subterm_def, ltree_paths_def, ltree_lookup] - >> UNBETA_TAC [subterm_alt] “subterm X M (h::p) r” - >> UNBETA_TAC [BT_def, Once ltree_unfold, BT_generator_def] “BT' X M r” - >> simp [LMAP_fromList, EL_MAP, Abbr ‘l’] - >> ‘n = n'’ by rw [Abbr ‘n’, Abbr ‘n'’] - >> POP_ASSUM (fs o wrap o SYM) - >> Q.PAT_X_ASSUM ‘vs = vs'’ (fs o wrap o SYM) - >> Q.PAT_X_ASSUM ‘M1 = M1'’ (fs o wrap o SYM) - >> Q.PAT_X_ASSUM ‘Ms = Ms'’ (fs o wrap o SYM) - >> qunabbrev_tac ‘vs’ + >- rw [subterm_def, BT_of_unsolvables, ltree_el_def] + (* stage work *) + >> rw [subterm_def, BT_def, BT_generator_def, Once ltree_unfold, ltree_el_def] + >> qabbrev_tac ‘M0 = principle_hnf M’ + >> qabbrev_tac ‘n = LAMl_size M0’ >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ >> ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] - (* extra work *) - >> qunabbrev_tac ‘y’ + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, “y :string”, “args :term list”)) ‘M1’ >> ‘TAKE n vs = vs’ by rw [] >> POP_ASSUM (rfs o wrap) + >> qabbrev_tac ‘m = LENGTH args’ + >> simp [LNTH_fromList, GSYM BT_def, EL_MAP] >> Cases_on ‘h < m’ >> simp [] - >> ‘Ms = args’ by rw [Abbr ‘Ms’] - >> POP_ASSUM (fs o wrap) - >> DISCH_TAC - >> simp [GSYM BT_def] - >> fs [ltree_paths_def, ltree_lookup_def, LNTH_fromList, EL_MAP] - >> T_TAC - (* extra work *) - >> FIRST_X_ASSUM MATCH_MP_TAC >> art [] + >> qabbrev_tac ‘N = EL h args’ + >> FIRST_X_ASSUM MATCH_MP_TAC + >> qunabbrev_tac ‘N’ >> MATCH_MP_TAC subterm_induction_lemma' >> qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] QED +Theorem ltree_paths_valid_thm : + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r ==> + (p IN ltree_paths (BT' X M r) <=> subterm X M p r <> NONE) +Proof + rw [ltree_paths_alt, BT_ltree_el_eq_none] +QED + +Theorem subterm_imp_ltree_paths : + !p X M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm X M p r <> NONE ==> + p IN ltree_paths (BT' X M r) +Proof + PROVE_TAC [ltree_paths_valid_thm] +QED + +(* p <> [] is required because ‘[] IN ltree_paths (BT' X M r)’ always holds. *) +Theorem ltree_paths_imp_solvable : + !p X M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ p <> [] /\ + p IN ltree_paths (BT' X M r) ==> solvable M +Proof + rw [ltree_paths_def] + >> Cases_on ‘p’ >> fs [] + >> CCONTR_TAC + >> fs [BT_of_unsolvables, ltree_lookup_def] +QED + (* ltree_lookup returns more information (the entire subtree), thus can be used to construct the return value of ltree_el. @@ -1120,14 +1171,6 @@ Proof >> qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] QED -Theorem subterm_is_none_exclusive : - !p X M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - p IN ltree_paths (BT' X M r) /\ - subterm X M p r = NONE ==> subterm X M (FRONT p) r <> NONE -Proof - METIS_TAC [subterm_is_none_iff_parent_unsolvable] -QED - (* NOTE: for whatever reasons such that ‘subterm X M p = NONE’, even when ‘p NOTIN ltree_paths (BT X M)’, the conclusion (rhs) always holds. *) @@ -1265,12 +1308,7 @@ Theorem subterm_valid_path_lemma : !q. q <<= FRONT p ==> subterm X M q r <> NONE Proof rpt GEN_TAC >> STRIP_TAC - >> Cases_on ‘subterm X M p r = NONE’ - >- (POP_ASSUM MP_TAC \\ - rw [subterm_is_none_iff_parent_unsolvable] \\ - Cases_on ‘FRONT p = []’ >- fs [] \\ - MP_TAC (Q.SPECL [‘X’, ‘M’, ‘FRONT p’, ‘r’] subterm_solvable_lemma) \\ - rw []) + >> ‘subterm X M p r <> NONE’ by PROVE_TAC [ltree_paths_valid_thm] >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_solvable_lemma) >> rw [] >> FIRST_X_ASSUM MATCH_MP_TAC @@ -1757,7 +1795,7 @@ Proof >> qabbrev_tac ‘N' = tpm p1 N’ >> T_TAC (* finally, using IH in a bulk way *) >> FIRST_X_ASSUM MATCH_MP_TAC - (* extra goal #1 (easy) *) + (* extra goal no.1 (easy) *) >> CONJ_TAC >- (simp [Abbr ‘N'’, SUBSET_DEF, FV_tpm] \\ rpt STRIP_TAC \\ @@ -1806,7 +1844,7 @@ Proof DISJ2_TAC \\ qunabbrev_tac ‘p1’ \\ MATCH_MP_TAC MEM_lswapstr >> rw []) - (* extra goal #2 (hard) *) + (* extra goal no.2 (hard) *) >> CONJ_TAC >- (simp [Abbr ‘N'’, FV_tpm, SUBSET_DEF] \\ simp [GSYM lswapstr_append, GSYM REVERSE_APPEND] \\ @@ -2362,221 +2400,6 @@ Proof RW_TAC std_ss [BT_tpm_thm, ltree_paths_map_cong] QED -(*---------------------------------------------------------------------------* - * Equivalent terms - *---------------------------------------------------------------------------*) - -(* Definition 10.2.19 [1, p. 238] - - M = LAMl v1 (y @* Ms) @@ (MAP VAR v1) == y @* Ms' - N = LAMl v2 (y' @* Ns) @@ (MAP VAR v2) == y' @* Ns' - - LENGTH Ms = n /\ LENGTH Ns = n' - LENGTH Ms' = m /\ LENGTH Ns' = m' - - y = y' - n - m = n' - m' (possibly negative) <=> n + m' = n' + m (non-negative) - *) -Definition equivalent_def : - equivalent (M :term) (N :term) = - if solvable M /\ solvable N then - let M0 = principle_hnf M; - N0 = principle_hnf N; - n = LAMl_size M0; - n' = LAMl_size N0; - vs = NEWS (MAX n n') (FV M UNION FV N); - vsM = TAKE n vs; - vsN = TAKE n' vs; - M1 = principle_hnf (M0 @* MAP VAR vsM); - N1 = principle_hnf (N0 @* MAP VAR vsN); - y = hnf_head M1; - y' = hnf_head N1; - m = LENGTH (hnf_children M1); - m' = LENGTH (hnf_children N1); - in - y = y' /\ n + m' = n' + m - else - ~solvable M /\ ~solvable N -End - -Theorem equivalent_reflexive : - reflexive equivalent -Proof - rw [reflexive_def, equivalent_def] -QED - -(* |- equivalent x x *) -Theorem equivalent_refl[simp] = - SPEC_ALL (REWRITE_RULE [reflexive_def] equivalent_reflexive) - -Theorem equivalent_symmetric : - symmetric equivalent -Proof - RW_TAC std_ss [symmetric_def, equivalent_def, Once MAX_COMM, Once UNION_COMM] - >> reverse (Cases_on ‘solvable x /\ solvable y’) >- fs [] - >> simp [] - >> rename1 ‘y1 = y2 /\ m1 + n = m + n1 <=> y3 = y4 /\ m3 + n1 = m2 + n3’ - >> ‘n3 = n’ by rw [Abbr ‘n3’, Abbr ‘n’] >> gs [] - >> EQ_TAC >> rw [] -QED - -(* |- !x y. equivalent x y <=> equivalent y x *) -Theorem equivalent_comm = REWRITE_RULE [symmetric_def] equivalent_symmetric - -Theorem equivalent_of_solvables : - !M N. solvable M /\ solvable N ==> - (equivalent M N <=> - let M0 = principle_hnf M; - N0 = principle_hnf N; - n = LAMl_size M0; - n' = LAMl_size N0; - vs = NEWS (MAX n n') (FV M UNION FV N); - vsM = TAKE n vs; - vsN = TAKE n' vs; - M1 = principle_hnf (M0 @* MAP VAR vsM); - N1 = principle_hnf (N0 @* MAP VAR vsN); - y = hnf_head M1; - y' = hnf_head N1; - m = LENGTH (hnf_children M1); - m' = LENGTH (hnf_children N1); - in - y = y' /\ n + m' = n' + m) -Proof - RW_TAC std_ss [equivalent_def] -QED - -(* beta-equivalent terms are also equivalent here *) -Theorem lameq_imp_equivalent : - !M N. M == N ==> equivalent M N -Proof - rpt STRIP_TAC - >> reverse (Cases_on ‘solvable M’) - >- (‘unsolvable N’ by METIS_TAC [lameq_solvable_cong] \\ - rw [equivalent_def]) - >> ‘solvable N’ by METIS_TAC [lameq_solvable_cong] - >> qabbrev_tac ‘X = FV M UNION FV N’ - >> ‘FINITE X’ by rw [Abbr ‘X’] - >> ‘LAMl_size (principle_hnf M) = LAMl_size (principle_hnf N)’ - by METIS_TAC [lameq_principle_hnf_size_eq'] - (* stage work *) - >> RW_TAC std_ss [equivalent_of_solvables] (* 2 subgoals, same tactics *) - >> ‘ALL_DISTINCT vs /\ DISJOINT (set vs) X /\ LENGTH vs = n’ - by (rw [Abbr ‘vs’, NEWS_def]) - >> ‘vsM = vs’ by rw [Abbr ‘vsM’, TAKE_LENGTH_ID_rwt] - >> POP_ASSUM (fs o wrap) - >> Q.PAT_X_ASSUM ‘vs = vsN’ (fs o wrap o SYM) - >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘N’, ‘M0’, ‘N0’, ‘n’, ‘vs’, ‘M1’, ‘N1’] - lameq_principle_hnf_thm_simple) - >> simp [Abbr ‘X’, GSYM solvable_iff_has_hnf] -QED - -(* NOTE: the initial calls of ‘principle_hnf’ get eliminated if the involved - terms are already in head normal forms. - *) -Theorem equivalent_of_hnf : - !M N. hnf M /\ hnf N ==> - (equivalent M N <=> - let n = LAMl_size M; - n' = LAMl_size N; - vs = NEWS (MAX n n') (FV M UNION FV N); - vsM = TAKE n vs; - vsN = TAKE n' vs; - M1 = principle_hnf (M @* MAP VAR vsM); - N1 = principle_hnf (N @* MAP VAR vsN); - y = hnf_head M1; - y' = hnf_head N1; - m = LENGTH (hnf_children M1); - m' = LENGTH (hnf_children N1) - in - y = y' /\ n + m' = n' + m) -Proof - rpt STRIP_TAC - >> ‘solvable M /\ solvable N’ by PROVE_TAC [hnf_has_hnf, solvable_iff_has_hnf] - >> RW_TAC std_ss [equivalent_def, principle_hnf_reduce] - >> METIS_TAC [] -QED - -(* From [1, p.238]. This concerte example shows that dB encoding is not easy in - defining this "concept": the literal encoding of inner head variables are not - the same for equivalent terms. - *) -Theorem not_equivalent_example : - !x y. x <> y ==> ~equivalent (LAM x (VAR y @@ t)) (LAM y (VAR y @@ t)) -Proof - qx_genl_tac [‘x’, ‘v’] >> DISCH_TAC - >> ‘hnf (LAM x (VAR v @@ t)) /\ hnf (LAM v (VAR v @@ t))’ by rw [] - >> ‘solvable (LAM x (VAR v @@ t)) /\ solvable (LAM v (VAR v @@ t))’ - by rw [solvable_iff_has_hnf, hnf_has_hnf] - >> RW_TAC std_ss [equivalent_of_solvables, principle_hnf_reduce] - (* fix M0 *) - >> qunabbrev_tac ‘M0’ >> qabbrev_tac ‘M0 = LAM x (VAR v @@ t)’ - >> ‘ALL_DISTINCT vs /\ DISJOINT (set vs) (FV M0 UNION FV N0) /\ - LENGTH vs = MAX n n'’ by rw [Abbr ‘vs’, NEWS_def] - >> ‘DISJOINT (set vs) (FV M0) /\ DISJOINT (set vs) (FV N0)’ - by METIS_TAC [DISJOINT_SYM, DISJOINT_UNION] - >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, - “y1 :string”, “args1 :term list”)) ‘M1’ - >> Q_TAC (HNF_TAC (“N0 :term”, “vs :string list”, - “y2 :string”, “args2 :term list”)) ‘N1’ - >> ‘TAKE (LAMl_size M0) vs = vsM’ by rw [Abbr ‘vsM’, Abbr ‘n’] - >> ‘TAKE (LAMl_size N0) vs = vsN’ by rw [Abbr ‘vsN’, Abbr ‘n'’] - >> NTAC 2 (POP_ASSUM (rfs o wrap)) - (* reshaping and reordering assumptions *) - >> qunabbrev_tac ‘M1’ - >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vsM)’ - >> qunabbrev_tac ‘N1’ - >> qabbrev_tac ‘N1 = principle_hnf (N0 @* MAP VAR vsN)’ - >> Q.PAT_X_ASSUM ‘M0 = _’ ASSUME_TAC - >> Q.PAT_X_ASSUM ‘N0 = _’ ASSUME_TAC - >> Q.PAT_X_ASSUM ‘M1 = _’ ASSUME_TAC - >> Q.PAT_X_ASSUM ‘N1 = _’ ASSUME_TAC - >> ‘VAR y1 = y’ by rw [Abbr ‘y’ , absfree_hnf_head] - >> ‘VAR y2 = y'’ by rw [Abbr ‘y'’, absfree_hnf_head] - >> qunabbrevl_tac [‘n’, ‘n'’] - >> Know ‘LAMl_size M0 = 1 /\ LAMl_size N0 = 1’ - >- (rw [Abbr ‘M0’, Abbr ‘N0’, LAMl_size_def]) - >> DISCH_THEN (rfs o wrap) - >> ‘vsN = vs’ by rw [Abbr ‘vsN’, TAKE_LENGTH_ID_rwt] - >> POP_ASSUM (rfs o wrap) - >> Q.PAT_X_ASSUM ‘vs = vsM’ (rfs o wrap o SYM) - >> qunabbrev_tac ‘vsN’ - (* stage work *) - >> qabbrev_tac ‘z = HD vs’ - >> ‘vs = [z]’ by METIS_TAC [SING_HD] - >> POP_ASSUM (rfs o wrap) - >> qunabbrevl_tac [‘M0’, ‘N0’] - >> DISJ1_TAC - >> qunabbrevl_tac [‘y’, ‘y'’] - >> Q.PAT_X_ASSUM ‘VAR y1 = hnf_head M1’ (rfs o wrap o SYM) - >> Q.PAT_X_ASSUM ‘_ = LAM z (VAR y1 @* args1)’ (rfs o wrap o SYM) - >> Q.PAT_X_ASSUM ‘_ = LAM z (VAR y2 @* args2)’ (rfs o wrap o SYM) - (* now the goal is ‘y1 <> y2’ *) - >> qabbrev_tac ‘u = VAR v @@ t’ - >> ‘hnf u’ by rw [Abbr ‘u’] - >> Know ‘M1 = [VAR z/x] u’ - >- (qunabbrev_tac ‘M1’ \\ - Cases_on ‘z = x’ >- (POP_ASSUM (gs o wrap) \\ - fs [principle_hnf_beta_reduce1]) \\ - MATCH_MP_TAC principle_hnf_beta >> simp [Abbr ‘u’] \\ - rfs [FV_thm]) - >> DISCH_THEN (rfs o wrap) - >> Know ‘N1 = [VAR z/v] u’ - >- (qunabbrev_tac ‘N1’ \\ - Cases_on ‘z = v’ >- (POP_ASSUM (rfs o wrap)) \\ - MATCH_MP_TAC principle_hnf_beta >> simp [Abbr ‘u’] \\ - rfs [FV_thm]) - >> DISCH_THEN (rfs o wrap) - >> qunabbrevl_tac [‘M1’, ‘N1’] - >> rfs [Abbr ‘u’, app_eq_appstar] - >> METIS_TAC [] -QED - -Theorem equivalent_of_unsolvables : - !M N. unsolvable M /\ unsolvable N ==> equivalent M N -Proof - rw [equivalent_def] -QED - (*---------------------------------------------------------------------------* * Boehm transformations *---------------------------------------------------------------------------*) @@ -2652,23 +2475,21 @@ QED (* ‘apply pi M’ (applying a Boehm transformation) means "M^{pi}" or "pi(M)" NOTE: ‘apply [f3;f2;f1] M = (f3 o f2 o f1) M = f3 (f2 (f1 M))’ - - NOTE2: This function seems general: “:('a -> 'a) list -> 'a -> 'a”. *) -Definition apply_def : - apply_functions pi = FOLDR $o I pi +Definition apply_transform_def : + apply_transform (pi :transform) = FOLDR $o I pi End -Overload apply = “apply_functions” +Overload apply = “apply_transform” (* NOTE: either FOLDL or FOLDR is correct (due to combinTheory.o_ASSOC), but FOLDR seems more natural requiring natural list induction in the next proof(s), while FOLDL would require SNOC_INDUCT. *) -Theorem apply_alt : +Theorem apply_transform_alt : !pi. apply pi = FOLDL $o I pi Proof - REWRITE_TAC [apply_def] + REWRITE_TAC [apply_transform_def] >> LIST_INDUCT_TAC >> rw [FOLDL, FOLDR] >> KILL_TAC (* only FOLDL is left *) >> qid_spec_tac ‘pi’ >> SNOC_INDUCT_TAC @@ -2676,13 +2497,13 @@ Proof >> POP_ASSUM (rw o wrap o SYM) QED -Theorem apply_rwts[simp] : +Theorem apply_transform_rwts[simp] : (apply [] = I) /\ (!f pi M. apply (f::pi) M = f (apply pi M)) /\ (!f pi M. apply (SNOC f pi) M = apply pi (f M)) Proof - NTAC 2 (CONJ_TAC >- rw [apply_def, o_DEF]) - >> rw [apply_alt, o_DEF, FOLDL_SNOC] + NTAC 2 (CONJ_TAC >- rw [apply_transform_def, o_DEF]) + >> rw [apply_transform_alt, o_DEF, FOLDL_SNOC] QED (* Lemma 10.3.4 (i) [1, p.246] *) @@ -2690,9 +2511,9 @@ Theorem Boehm_transform_lameq_ctxt : !pi. Boehm_transform pi ==> ?c. ctxt c /\ !M. apply pi M == c M Proof Induct_on ‘pi’ - >> rw [Boehm_transform_def, apply_def] + >> rw [Boehm_transform_def, apply_transform_def] >- (Q.EXISTS_TAC ‘\x. x’ >> rw [ctxt_rules, FOLDR]) - >> fs [GSYM Boehm_transform_def, apply_def] + >> fs [GSYM Boehm_transform_def, apply_transform_def] >> fs [solving_transform_def] >- (Q.EXISTS_TAC ‘\y. c y @@ (\y. VAR x) y’ \\ rw [ctxt_rules, FOLDR] \\ @@ -2804,7 +2625,7 @@ QED Theorem Boehm_apply_SNOC_SUB : !(N :term) v p M. apply (SNOC [N/v] p) M = apply p ([N/v] M) Proof - rw [apply_def, FOLDR_SNOC] + rw [apply_transform_def, FOLDR_SNOC] QED Theorem Boehm_apply_MAP_rightctxt : @@ -2823,19 +2644,23 @@ Proof >> rw [Boehm_apply_MAP_rightctxt] QED -(* NOTE: if M is solvable, ‘apply pi M’ may not be solvable. *) -Theorem Boehm_apply_unsolvable : +(* |- !M N. solvable (M @@ N) ==> solvable M *) +Theorem solvable_APP_E[local] = + has_hnf_APP_E |> REWRITE_RULE [GSYM solvable_iff_has_hnf] + |> Q.GENL [‘M’, ‘N’] + +Theorem Boehm_transform_of_unsolvables : !pi M. Boehm_transform pi /\ unsolvable M ==> unsolvable (apply pi M) Proof - SNOC_INDUCT_TAC >> rw [] - >> FIRST_X_ASSUM MATCH_MP_TAC >> art [] - >> fs [solving_transform_def, solvable_iff_has_hnf] (* 2 subgaols *) - >| [ (* goal 1 (of 2) *) - CCONTR_TAC >> fs [] \\ - METIS_TAC [has_hnf_APP_E], - (* goal 2 (of 2) *) - CCONTR_TAC >> fs [] \\ - METIS_TAC [has_hnf_SUB_E] ] + Induct_on ‘pi’ using SNOC_INDUCT + >- rw [] + >> simp [FOLDR_SNOC, Boehm_transform_SNOC] + >> qx_genl_tac [‘t’, ‘M’] + >> reverse (rw [solving_transform_def]) + >- (FIRST_X_ASSUM MATCH_MP_TAC >> art [] \\ + MATCH_MP_TAC unsolvable_subst >> art []) + >> FIRST_X_ASSUM MATCH_MP_TAC >> simp [] + >> PROVE_TAC [solvable_APP_E] QED (* Definition 10.3.5 (ii) *) @@ -2859,10 +2684,6 @@ Definition is_ready_def : ?N. M -h->* N /\ hnf N /\ ~is_abs N /\ head_original N End -Definition is_ready' : - is_ready' M <=> is_ready M /\ solvable M -End - (* NOTE: This alternative definition of ‘is_ready’ consumes ‘head_original’ and eliminated the ‘principle_hnf’ inside it. *) @@ -2897,15 +2718,6 @@ Proof >> qexistsl_tac [‘y’, ‘args’] >> art [] QED -Theorem is_ready_alt' : - !M. is_ready' M <=> solvable M /\ - ?y Ns. M -h->* VAR y @* Ns /\ EVERY (\e. y # e) Ns -Proof - (* NOTE: this proof relies on the new [NOT_AND'] in boolSimps.BOOL_ss *) - rw [is_ready', is_ready_alt, RIGHT_AND_OVER_OR] - >> REWRITE_TAC [Once CONJ_COMM] -QED - (* ‘subterm_width M p’ is the maximal number of children of all subterms of form ‘subterm X M q r’ such that ‘q <<= p’, irrelevant with particular X and r. @@ -2913,10 +2725,10 @@ QED *) Definition subterm_width_def[nocompute] : subterm_width M p = - MAX_SET (IMAGE (\q. let N = subterm (FV M) M q 0 in - if N <> NONE /\ solvable (FST (THE N)) then - hnf_children_size (principle_hnf (FST (THE N))) - else 0) {q | q <<= p}) + MAX_SET (IMAGE (\q. let N = subterm (FV M) M q 0 in + if N <> NONE /\ solvable (FST (THE N)) then + hnf_children_size (principle_hnf (FST (THE N))) + else 0) {q | q <<= p}) End (* ‘subterm_length M p’ is the maximal length of LAMl binding list of subterms @@ -2926,10 +2738,10 @@ End *) Definition subterm_length_def[nocompute] : subterm_length M p = - MAX_SET (IMAGE (\q. let N = subterm (FV M) M q 0 in - if N <> NONE /\ solvable (FST (THE N)) then - LAMl_size (principle_hnf (FST (THE N))) - else 0) {q | q <<= p}) + MAX_SET (IMAGE (\q. let N = subterm (FV M) M q 0 in + if N <> NONE /\ solvable (FST (THE N)) then + LAMl_size (principle_hnf (FST (THE N))) + else 0) {q | q <<= p}) End Theorem subterm_width_nil : @@ -2990,14 +2802,12 @@ QED Theorem subterm_width_thm : !X M p r. - FINITE X /\ FV M SUBSET X UNION RANK r ==> - subterm_width M p = - MAX_SET - (IMAGE - (\q. let N = subterm X M q r in - if N <> NONE /\ solvable (FST (THE N)) then - hnf_children_size (principle_hnf (FST (THE N))) - else 0) {q | q <<= p}) + FINITE X /\ FV M SUBSET X UNION RANK r ==> + subterm_width M p = + MAX_SET (IMAGE (\q. let N = subterm X M q r in + if N <> NONE /\ solvable (FST (THE N)) then + hnf_children_size (principle_hnf (FST (THE N))) + else 0) {q | q <<= p}) Proof rw [subterm_width_def] >> AP_TERM_TAC @@ -3017,14 +2827,12 @@ QED Theorem subterm_length_thm : !X M p r. - FINITE X /\ FV M SUBSET X UNION RANK r ==> - subterm_length M p = - MAX_SET - (IMAGE - (\q. let N = subterm X M q r in - if N <> NONE /\ solvable (FST (THE N)) then - LAMl_size (principle_hnf (FST (THE N))) - else 0) {q | q <<= p}) + FINITE X /\ FV M SUBSET X UNION RANK r ==> + subterm_length M p = + MAX_SET (IMAGE (\q. let N = subterm X M q r in + if N <> NONE /\ solvable (FST (THE N)) then + LAMl_size (principle_hnf (FST (THE N))) + else 0) {q | q <<= p}) Proof rw [subterm_length_def] >> AP_TERM_TAC @@ -3043,8 +2851,8 @@ Proof QED Theorem subterm_width_first : - !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M ==> - hnf_children_size (principle_hnf M) <= subterm_width M p + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M + ==> hnf_children_size (principle_hnf M) <= subterm_width M p Proof rpt STRIP_TAC >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_thm) @@ -3057,8 +2865,8 @@ Proof QED Theorem subterm_length_first : - !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M ==> - LAMl_size (principle_hnf M) <= subterm_length M p + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M + ==> LAMl_size (principle_hnf M) <= subterm_length M p Proof rpt STRIP_TAC >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_length_thm) @@ -3071,11 +2879,11 @@ Proof QED Theorem subterm_width_last : - !X M p q r. FINITE X /\ FV M SUBSET X UNION RANK r /\ q <<= p /\ - subterm X M q r <> NONE /\ - solvable (subterm' X M q r) ==> - hnf_children_size (principle_hnf (subterm' X M q r)) <= - subterm_width M p + !X M p q r. + FINITE X /\ FV M SUBSET X UNION RANK r /\ q <<= p /\ + subterm X M q r <> NONE /\ + solvable (subterm' X M q r) + ==> hnf_children_size (principle_hnf (subterm' X M q r)) <= subterm_width M p Proof rpt STRIP_TAC >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_thm) @@ -3088,11 +2896,11 @@ Proof QED Theorem subterm_length_last : - !X M p q r. FINITE X /\ FV M SUBSET X UNION RANK r /\ q <<= p /\ - subterm X M q r <> NONE /\ - solvable (subterm' X M q r) ==> - LAMl_size (principle_hnf (subterm' X M q r)) <= - subterm_length M p + !X M p q r. + FINITE X /\ FV M SUBSET X UNION RANK r /\ q <<= p /\ + subterm X M q r <> NONE /\ + solvable (subterm' X M q r) + ==> LAMl_size (principle_hnf (subterm' X M q r)) <= subterm_length M p Proof rpt STRIP_TAC >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_length_thm) @@ -3105,11 +2913,14 @@ Proof QED Theorem solvable_subst_permutator : - !X M r P v d. FINITE X /\ FV M SUBSET X UNION RANK r /\ v IN X UNION RANK r /\ - P = permutator d /\ hnf_children_size (principle_hnf M) <= d /\ - solvable M ==> solvable ([P/v] M) + !X M r P v d. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + v IN X UNION RANK r /\ P = permutator d /\ + solvable M /\ subterm_width M [] <= d + ==> solvable ([P/v] M) /\ subterm_width ([P/v] M) [] <= d Proof - RW_TAC std_ss [] + rpt GEN_TAC >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘P = permutator d’ (REWRITE_TAC o wrap) >> qabbrev_tac ‘P = permutator d’ >> qabbrev_tac ‘M0 = principle_hnf M’ >> qabbrev_tac ‘n = LAMl_size M0’ @@ -3122,7 +2933,6 @@ Proof >> POP_ASSUM (rfs o wrap) >> ‘M0 == M’ by rw [Abbr ‘M0’, lameq_principle_hnf'] >> ‘[P/v] M0 == [P/v] M’ by rw [lameq_sub_cong] - >> Suff ‘solvable ([P/v] M0)’ >- PROVE_TAC [lameq_solvable_cong] >> ‘FV P = {}’ by rw [Abbr ‘P’, FV_permutator] >> ‘DISJOINT (set vs) (FV P)’ by rw [DISJOINT_ALT'] >> Know ‘~MEM v vs’ @@ -3134,133 +2944,179 @@ Proof qunabbrev_tac ‘vs’ \\ MATCH_MP_TAC DISJOINT_RANK_RNEWS' >> art []) >> DISCH_TAC + >> CONJ_ASM1_TAC + >- (Suff ‘solvable ([P/v] M0)’ >- PROVE_TAC [lameq_solvable_cong] \\ + simp [LAMl_SUB, appstar_SUB] \\ + reverse (Cases_on ‘y = v’) + >- (simp [SUB_THM, solvable_iff_has_hnf] \\ + MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ + simp [solvable_iff_has_hnf, has_hnf_thm] \\ + qabbrev_tac ‘args' = MAP [P/v] args’ \\ + qabbrev_tac ‘m = LENGTH args’ \\ + ‘LENGTH args' = m’ by rw [Abbr ‘args'’] \\ + Q.PAT_X_ASSUM ‘subterm_width M [] <= d’ MP_TAC \\ + rw [hnf_children_size_thm, subterm_width_nil] \\ + (* applying hreduce_permutator_thm *) + MP_TAC (Q.SPECL [‘{}’, ‘d’, ‘args'’] hreduce_permutator_thm) \\ + rw [Abbr ‘P’] \\ + Q.EXISTS_TAC ‘LAMl xs (LAM y (VAR y @* args' @* MAP VAR xs))’ \\ + rw [hnf_appstar, hnf_thm]) + (* extra goal for induction *) + >> Q.PAT_X_ASSUM ‘subterm_width M [] <= d’ MP_TAC + >> simp [subterm_width_nil] + >> DISCH_TAC + (* applying principle_hnf_hreduce, hreduces_hnf_imp_principle_hnf, etc. + + M -h->* M0 = LAMl vs (VAR y @* args) + [P/v] M -h->* [P/v] (LAMl vs (VAR y @* args)) + *) + >> Know ‘[P/v] M -h->* [P/v] M0’ + >- (MATCH_MP_TAC hreduce_substitutive \\ + METIS_TAC [principle_hnf_thm']) >> simp [LAMl_SUB, appstar_SUB] >> reverse (Cases_on ‘y = v’) >- (simp [SUB_THM, solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) - >> simp [solvable_iff_has_hnf, has_hnf_thm] + DISCH_TAC \\ + qabbrev_tac ‘args' = MAP [P/v] args’ \\ + ‘hnf (LAMl vs (VAR y @* args'))’ by rw [hnf_appstar] \\ + ‘principle_hnf ([P/v] M) = LAMl vs (VAR y @* args')’ + by METIS_TAC [principle_hnf_thm'] >> POP_ORW \\ + qabbrev_tac ‘m = LENGTH args’ \\ + ‘LENGTH args' = m’ by rw [Abbr ‘args'’] \\ + simp []) + (* stage work *) + >> simp [] >> qabbrev_tac ‘args' = MAP [P/v] args’ + >> DISCH_TAC >> qabbrev_tac ‘m = LENGTH args’ >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’] - >> fs [hnf_children_size_thm] - (* applying hreduce_permutator_thm *) >> MP_TAC (Q.SPECL [‘{}’, ‘d’, ‘args'’] hreduce_permutator_thm) - >> rw [Abbr ‘P’] - >> Q.EXISTS_TAC ‘LAMl xs (LAM y (VAR y @* args' @* MAP VAR xs))’ - >> rw [hnf_appstar, hnf_thm] + >> simp [] + >> STRIP_TAC + >> ‘LAMl vs (P @* args') -h->* + LAMl vs (LAMl xs (LAM y' (VAR y' @* args' @* MAP VAR xs)))’ + by rw [hreduce_LAMl] + >> Know ‘[P/v] M -h->* LAMl vs (LAMl xs (LAM y' (VAR y' @* args' @* MAP VAR xs)))’ + >- (MATCH_MP_TAC hreduce_TRANS \\ + Q.EXISTS_TAC ‘LAMl vs (P @* args')’ >> art []) + >> REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND, GSYM LAMl_SNOC] + >> qabbrev_tac ‘ys = SNOC y' (vs ++ xs)’ + >> qabbrev_tac ‘args2 = args' ++ MAP VAR xs’ + >> DISCH_TAC + >> ‘hnf (LAMl ys (VAR y' @* args2))’ by rw [hnf_appstar] + >> ‘principle_hnf ([P/v] M) = LAMl ys (VAR y' @* args2)’ + by METIS_TAC [principle_hnf_thm'] + >> POP_ORW + >> simp [Abbr ‘args2’] QED Theorem solvable_subst_permutator_cong : !X M r P v d. - FINITE X /\ FV M SUBSET X UNION RANK r /\ v IN X UNION RANK r /\ - P = permutator d /\ hnf_children_size (principle_hnf M) <= d - ==> (solvable ([P/v] M) <=> solvable M) + FINITE X /\ FV M SUBSET X UNION RANK r /\ + v IN X UNION RANK r /\ P = permutator d /\ + subterm_width M [] <= d ==> + (solvable ([P/v] M) <=> solvable M) Proof rpt STRIP_TAC >> EQ_TAC >- PROVE_TAC [unsolvable_subst] >> DISCH_TAC - >> MATCH_MP_TAC solvable_subst_permutator + >> MATCH_MP_TAC (cj 1 solvable_subst_permutator) >> qexistsl_tac [‘X’, ‘r’, ‘d’] >> art [] QED -Theorem solvable_ISUB_permutator : - !X M M0 r ss d. - FINITE X /\ FV M SUBSET X UNION RANK r /\ - M0 = principle_hnf M /\ hnf_children_size M0 <= d /\ - DOM ss SUBSET X UNION RANK r /\ - (!N. MEM N (MAP FST ss) ==> N = permutator d) /\ - solvable M ==> solvable (M ISUB ss) +Theorem solvable_isub_permutator : + !X r d ss M. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + DOM ss SUBSET X UNION RANK r /\ + (!P. MEM P (MAP FST ss) ==> P = permutator d) /\ + solvable M /\ subterm_width M [] <= d + ==> solvable (M ISUB ss) /\ + subterm_width (M ISUB ss) [] <= d Proof - RW_TAC std_ss [] - >> qabbrev_tac ‘P = permutator d’ - >> qabbrev_tac ‘M0 = principle_hnf M’ - >> qabbrev_tac ‘n = LAMl_size M0’ - >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ - >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ - >> ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] - >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, - “y :string”, “args :term list”)) ‘M1’ - >> ‘TAKE n vs = vs’ by rw [] - >> POP_ASSUM (rfs o wrap) - >> ‘M0 == M’ by rw [Abbr ‘M0’, lameq_principle_hnf'] - >> ‘M0 ISUB ss == M ISUB ss’ by rw [lameq_isub_cong] - >> Suff ‘solvable (M0 ISUB ss)’ >- PROVE_TAC [lameq_solvable_cong] - >> ‘FV P = {}’ by rw [Abbr ‘P’, FV_permutator] - >> ‘DISJOINT (set vs) (FV P)’ by rw [DISJOINT_ALT'] - >> Know ‘DISJOINT (set vs) (DOM ss)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘X UNION RANK r’ >> art [] \\ - simp [DISJOINT_UNION', Abbr ‘vs’] \\ - MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) - >> DISCH_TAC - >> Know ‘FVS ss = {}’ - >- (rw [FVS_ALT] \\ - Cases_on ‘ss = []’ >> rw [Once EXTENSION] \\ - Q.PAT_X_ASSUM ‘!N. MEM N (MAP FST ss) ==> N = P’ MP_TAC \\ - rw [o_DEF, MEM_MAP] \\ - EQ_TAC >> rw [] - >- (rename1 ‘MEM s ss’ >> PairCases_on ‘s’ \\ - Q.PAT_X_ASSUM ‘!N. _’ (MP_TAC o Q.SPEC ‘s0’) \\ - impl_tac >- (Q.EXISTS_TAC ‘(s0,s1)’ >> rw []) \\ - DISCH_THEN (fs o wrap)) \\ - MP_TAC (ISPEC “set (ss :(term # string) list)” MEMBER_NOT_EMPTY) \\ - rw [] \\ - Q.EXISTS_TAC ‘x’ >> art [] \\ - PairCases_on ‘x’ \\ - Q.PAT_X_ASSUM ‘!N. _’ (MP_TAC o Q.SPEC ‘x0’) \\ - impl_tac >- (Q.EXISTS_TAC ‘(x0,x1)’ >> rw []) \\ - DISCH_THEN (fs o wrap)) - >> DISCH_TAC - >> ‘DISJOINT (set vs) (FVS ss)’ by rw [] - >> simp [LAMl_ISUB, appstar_ISUB] - >> reverse (Cases_on ‘y IN DOM ss’) - >- (simp [ISUB_VAR_FRESH', solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf \\ - rw [hnf_appstar]) - >> simp [solvable_iff_has_hnf, has_hnf_thm] - >> qabbrev_tac ‘args' = MAP (\t. t ISUB ss) args’ - >> qabbrev_tac ‘m = LENGTH args’ - >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’] - >> Know ‘VAR y ISUB ss = P’ - >- (Q.PAT_X_ASSUM ‘y IN DOM ss’ MP_TAC \\ - Q.PAT_X_ASSUM ‘!N. MEM N (MAP FST ss) ==> N = P’ MP_TAC \\ - Q.PAT_X_ASSUM ‘FV P = {}’ MP_TAC \\ - KILL_TAC \\ - Induct_on ‘ss’ >- simp [DOM_DEF] \\ - simp [FORALL_PROD] >> qx_genl_tac [‘N’, ‘v’] \\ - rw [DOM_DEF, IN_UNION] - >- (simp [] \\ - MATCH_MP_TAC ISUB_unchanged >> rw []) \\ - gs [] \\ - Cases_on ‘v = y’ >> simp [] \\ - MATCH_MP_TAC ISUB_unchanged >> rw []) - >> DISCH_TAC - (* applying hreduce_permutator_thm *) - >> fs [hnf_children_size_thm] (* m <= d *) - >> MP_TAC (Q.SPECL [‘{}’, ‘d’, ‘args'’] hreduce_permutator_thm) - >> rw [Abbr ‘P’] - >> Q.EXISTS_TAC ‘LAMl xs (LAM y' (VAR y' @* args' @* MAP VAR xs))’ - >> rw [hnf_appstar, hnf_thm] + NTAC 3 GEN_TAC + >> Induct_on ‘ss’ >- rw [] + >> simp [FORALL_PROD, DOM_DEF] + >> qx_genl_tac [‘P’, ‘v’, ‘M’] + >> STRIP_TAC + >> qabbrev_tac ‘Q = permutator d’ + >> qabbrev_tac ‘N = [Q/v] M’ + >> FIRST_X_ASSUM MATCH_MP_TAC >> simp [] + >> CONJ_TAC + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] \\ + simp [Abbr ‘N’, FV_SUB, Abbr ‘Q’, FV_permutator] \\ + SET_TAC []) + >> qunabbrev_tac ‘N’ + >> MATCH_MP_TAC solvable_subst_permutator + >> qexistsl_tac [‘X’, ‘r’] >> simp [] QED -Theorem solvable_ISUB_permutator_cong : - !X M M0 r ss d. - FINITE X /\ FV M SUBSET X UNION RANK r /\ - M0 = principle_hnf M /\ hnf_children_size M0 <= d /\ - DOM ss SUBSET X UNION RANK r /\ - (!N. MEM N (MAP FST ss) ==> N = permutator d) ==> - (solvable (M ISUB ss) <=> solvable M) +Theorem solvable_isub_permutator_cong : + !X M r ss d. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm_width M [] <= d /\ + DOM ss SUBSET X UNION RANK r /\ + (!P. MEM P (MAP FST ss) ==> P = permutator d) ==> + (solvable (M ISUB ss) <=> solvable M) Proof rpt STRIP_TAC >> EQ_TAC >- METIS_TAC [unsolvable_ISUB] >> DISCH_TAC - >> MATCH_MP_TAC solvable_ISUB_permutator - >> qexistsl_tac [‘X’, ‘M0’, ‘r’, ‘d’] >> art [] + >> MATCH_MP_TAC (cj 1 solvable_isub_permutator) + >> qexistsl_tac [‘X’, ‘r’, ‘d’] >> art [] QED -Theorem subterm_width_induction_lemma : - !X M h p r M0 n n' m vs' M1 Ms d. - FINITE X /\ FV M SUBSET X UNION RANK r /\ +(* NOTE: This is the first of a series of lemmas of increasing permutators. + In theory the proof works if each permutator is larger than the previous + one, not necessary increased just by one (which is enough for now). + + In other words, ‘ss = GENLIST (\i. (permutator (d + i),y i)) k’ can be + replaced by ‘ss = GENLIST (\i. (permutator (f i),y i)) k’ where ‘f’ is + increasing (or non-decreasing). + *) +Theorem solvable_isub_permutator_alt : + !X r d y k ss M. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + (!i. i < k ==> y i IN X UNION RANK r) /\ + ss = GENLIST (\i. (permutator (d + i),y i)) k /\ + solvable M /\ subterm_width M [] <= d + ==> solvable (M ISUB ss) /\ + subterm_width (M ISUB ss) [] <= d + k +Proof + NTAC 4 GEN_TAC + >> Induct_on ‘k’ >- rw [] + >> qx_genl_tac [‘ss'’, ‘M’] + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘ss' = _’ (REWRITE_TAC o wrap) + >> SIMP_TAC std_ss [GENLIST, ISUB_SNOC] + >> qabbrev_tac ‘P = \i. permutator (d + i)’ >> fs [] + >> qabbrev_tac ‘ss = GENLIST (\i. (P i,y i)) k’ + >> Q.PAT_X_ASSUM ‘!M'. FV M' SUBSET X UNION RANK r /\ _ ==> _’ + (MP_TAC o Q.SPEC ‘M’) + >> simp [] + >> STRIP_TAC + >> qabbrev_tac ‘N = M ISUB ss’ + >> qabbrev_tac ‘Q = P k’ + >> qabbrev_tac ‘v = y k’ + >> qabbrev_tac ‘w = d + k’ + >> MP_TAC (Q.SPECL [‘X’, ‘N’, ‘r’, ‘Q’, ‘v’, ‘w’] solvable_subst_permutator) + >> simp [Abbr ‘Q’, Abbr ‘v’, Abbr ‘w’] + >> impl_tac + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] \\ + qunabbrev_tac ‘N’ \\ + MP_TAC (Q.SPECL [‘ss’, ‘M’] FV_ISUB_upperbound) \\ + Suff ‘FVS ss = {}’ >- simp [] \\ + simp [Abbr ‘ss’, FVS_ALT] \\ + Cases_on ‘k = 0’ >> simp [] \\ + DISJ2_TAC \\ + simp [MAP_GENLIST, LIST_TO_SET_GENLIST] \\ + simp [Abbr ‘P’, FV_permutator, o_DEF] \\ + simp [IMAGE_CONST]) + >> rw [] +QED + +Theorem subterm_width_induction_lemma : + !X M h p r M0 n n' m vs' M1 Ms d. + FINITE X /\ FV M SUBSET X UNION RANK r /\ solvable M /\ M0 = principle_hnf M /\ n = LAMl_size M0 /\ n <= n' /\ @@ -3691,9 +3547,7 @@ Proof >> Induct_on ‘q’ >- (rw [] \\ qabbrev_tac ‘P = permutator d’ \\ - reverse (Cases_on ‘solvable ([P/v] M)’) - >- rw [subterm_width_nil] \\ - rw [subterm_width_def] \\ + Cases_on ‘solvable ([P/v] M)’ >> rw [subterm_width_def] \\ ‘solvable M’ by PROVE_TAC [unsolvable_subst] \\ qabbrev_tac ‘M0 = principle_hnf M’ \\ qabbrev_tac ‘n = LAMl_size M0’ \\ @@ -3746,10 +3600,10 @@ Proof ‘hnf (LAMl vs (LAMl xs (LAM y (VAR y @* args' @* MAP VAR xs))))’ by rw [hnf_LAMl, hnf_appstar] \\ qabbrev_tac ‘P = permutator d’ \\ - ‘principle_hnf ([P/v] M) = - LAMl vs (LAMl xs (LAM y (VAR y @* args' @* MAP VAR xs)))’ - by METIS_TAC [principle_hnf_thm'] >> POP_ORW \\ - simp [hnf_children_size_LAMl, GSYM appstar_APPEND]) + ‘principle_hnf ([P/v] M) = + LAMl vs (LAMl xs (LAM y (VAR y @* args' @* MAP VAR xs)))’ + by METIS_TAC [principle_hnf_thm'] >> POP_ORW \\ + simp [hnf_children_size_LAMl, GSYM appstar_APPEND]) (* stage work *) >> rpt GEN_TAC >> STRIP_TAC >> ‘p IN ltree_paths (BT' X M r)’ by PROVE_TAC [subterm_imp_ltree_paths] @@ -3828,8 +3682,8 @@ Proof already a hnf (v <> y), or can be head-reduced to a hnf (v = y). *) >> Know ‘solvable ([P/v] M)’ - >- (MATCH_MP_TAC solvable_subst_permutator \\ - qexistsl_tac [‘X’, ‘r’, ‘d’] >> simp []) + >- (MATCH_MP_TAC (cj 1 solvable_subst_permutator) \\ + qexistsl_tac [‘X’, ‘r’, ‘d’] >> simp [subterm_width_nil]) >> DISCH_TAC (* Now we need to know the exact form of ‘principle_hnf ([P/v] M)’. @@ -3866,7 +3720,7 @@ Proof >> CONV_TAC (UNBETA_CONV “subterm X ([P/v] M) (h::q) r”) >> qmatch_abbrev_tac ‘f _’ >> ASM_SIMP_TAC std_ss [subterm_of_solvables] - >> BasicProvers.LET_ELIM_TAC + >> LET_ELIM_TAC >> Q.PAT_X_ASSUM ‘subterm X (EL h args) q (SUC r) <> NONE’ MP_TAC >> simp [Abbr ‘f’, hnf_children_hnf] >> DISCH_TAC (* subterm X (EL h args) q (SUC r) <> NONE *) @@ -4144,9 +3998,9 @@ Proof >> simp [] >> DISCH_THEN K_TAC (* applying SUBSET_MAX_SET *) >> MATCH_MP_TAC SUBSET_MAX_SET - >> CONJ_TAC (* FINITE #1 *) + >> CONJ_TAC (* FINITE _ *) >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) - >> CONJ_TAC (* FINITE #2 *) + >> CONJ_TAC (* FINITE _ *) >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) >> rw [SUBSET_DEF] (* this asserts q' <<= q *) >> Know ‘q' <<= t’ @@ -4231,21 +4085,64 @@ Proof >> rw [Abbr ‘P’, FV_permutator] QED -Theorem subterm_isub_permutator_cong'[local] : - !ys p X M r P d ss. +Theorem subterm_isub_permutator_cong_alt : + !X p r d y k ss M. FINITE X /\ FV M SUBSET X UNION RANK r /\ + (!i. i < k ==> y i IN X UNION RANK r) /\ + ss = GENLIST (\i. (permutator (d + i),y i)) k /\ subterm X M p r <> NONE /\ - P = permutator d /\ - subterm_width M p <= d /\ - (!y. MEM y ys ==> y IN X UNION RANK r) /\ - ss = MAP (\y. (P,y)) ys + subterm_width M p <= d ==> subterm X (M ISUB ss) p r <> NONE /\ + subterm_width (M ISUB ss) p <= d + k /\ subterm' X (M ISUB ss) p r = (subterm' X M p r) ISUB ss Proof - rpt GEN_TAC + qx_genl_tac [‘X’, ‘p’, ‘r’, ‘d’, ‘y’] + >> Induct_on ‘k’ >- rw [] + >> qx_genl_tac [‘ss'’, ‘M’] + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘ss' = _’ (REWRITE_TAC o wrap) + >> SIMP_TAC std_ss [GENLIST, ISUB_SNOC] + >> qabbrev_tac ‘P = \i. permutator (d + i)’ >> fs [] + >> qabbrev_tac ‘ss = GENLIST (\i. (P i,y i)) k’ + >> Q.PAT_X_ASSUM ‘!M'. FV M' SUBSET X UNION RANK r /\ _ ==> _’ + (MP_TAC o Q.SPEC ‘M’) + >> simp [] >> STRIP_TAC - >> MP_TAC (Q.SPECL [‘ys’, ‘p’, ‘X’, ‘M’, ‘r’, ‘P’, ‘d’, ‘ss’] - subterm_isub_permutator_cong) >> rw [] + >> qabbrev_tac ‘N = M ISUB ss’ + >> qabbrev_tac ‘Q = P k’ + >> qabbrev_tac ‘v = y k’ + >> qabbrev_tac ‘w = d + k’ + >> MP_TAC (Q.SPECL [‘p’, ‘X’, ‘N’, ‘r’, ‘v’, ‘Q’, ‘w’] + subterm_subst_permutator_cong) + >> simp [Abbr ‘Q’, Abbr ‘v’, Abbr ‘w’] + >> impl_tac + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] \\ + qunabbrev_tac ‘N’ \\ + MP_TAC (Q.SPECL [‘ss’, ‘M’] FV_ISUB_upperbound) \\ + Suff ‘FVS ss = {}’ >- simp [] \\ + simp [Abbr ‘ss’, FVS_ALT] \\ + Cases_on ‘k = 0’ >> simp [] \\ + DISJ2_TAC \\ + simp [MAP_GENLIST, LIST_TO_SET_GENLIST] \\ + simp [Abbr ‘P’, FV_permutator, o_DEF] \\ + simp [IMAGE_CONST]) + >> rw [] +QED + +Theorem subterm_isub_permutator_cong_alt' : + !X p r d y k ss M. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + (!i. i < k ==> y i IN X UNION RANK r) /\ + ss = GENLIST (\i. (permutator (d + i),y i)) k /\ + subterm X M p r <> NONE /\ + subterm_width M p <= d + ==> subterm X (M ISUB ss) p r <> NONE /\ + subterm' X (M ISUB ss) p r = (subterm' X M p r) ISUB ss +Proof + rpt GEN_TAC >> STRIP_TAC + >> MP_TAC (Q.SPECL [‘X’, ‘p’, ‘r’, ‘d’, ‘y’, ‘k’, ‘ss’, ‘M’] + subterm_isub_permutator_cong_alt) + >> simp [] QED (* This theorem is similar with subterm_subst_permutator_cong_lemma. P is @@ -4332,7 +4229,7 @@ Proof >> CONV_TAC (UNBETA_CONV “subterm X ([P/v] M) (h::q) r”) >> qmatch_abbrev_tac ‘f _’ >> ASM_SIMP_TAC std_ss [subterm_of_solvables] - >> BasicProvers.LET_ELIM_TAC + >> LET_ELIM_TAC >> Q.PAT_X_ASSUM ‘subterm X (EL h args) q (SUC r) <> NONE’ MP_TAC >> simp [Abbr ‘f’, hnf_children_hnf] >> DISCH_TAC (* subterm X (EL h args) q (SUC r) <> NONE *) @@ -4480,1324 +4377,1369 @@ Proof >> Q.EXISTS_TAC ‘p’ >> rw [] QED -(* This theorem is based on FV_subterm_lemma - - NOTE: If initially ‘FV M SUBSET X UNION RANK r’ and ‘v # M’ but - ‘v IN RANK r’, each deeper subterm only adds FV from ROW r onwards, - thus ‘v # subterm X M p r’ still holds. +(* NOTE: ‘ltree_paths (BT' X M r) SUBSET ltree_paths (BT' X ([P/v] M) r)’ doesn't + hold. Instead, we need to consider certain p and ‘d <= subterm_width M p’. + This theorem holds even when M is not solvable. *) -Theorem FV_subterm_thm : - !X v p M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - subterm X M p r <> NONE /\ - v # M /\ v IN X UNION RANK r ==> v # (subterm' X M p r) +Theorem subterm_width_subst_permutator_cong : + !X P d v p M r. + FINITE X /\ FV M SUBSET X UNION RANK r /\ v IN X UNION RANK r /\ + P = permutator d /\ subterm_width M p <= d /\ + p IN ltree_paths (BT' X M r) ==> + p IN ltree_paths (BT' X ([P/v] M) r) /\ + subterm_width ([P/v] M) p <= d Proof - NTAC 2 GEN_TAC - >> Induct_on ‘p’ >- rw [] - >> rpt GEN_TAC - >> STRIP_TAC - >> Know ‘(!q. q <<= h::p ==> subterm X M q r <> NONE) /\ - (!q. q <<= FRONT (h::p) ==> solvable (subterm' X M q r))’ - >- (MATCH_MP_TAC subterm_solvable_lemma >> simp []) - >> STRIP_TAC - >> Know ‘solvable M’ - >- (Q.PAT_X_ASSUM ‘!q. q <<= FRONT (h::p) ==> solvable _’ - (MP_TAC o (Q.SPEC ‘[]’)) >> rw []) + simp [ltree_paths_def] + >> NTAC 4 GEN_TAC + >> Induct_on ‘p’ + >- (rw [ltree_lookup] \\ + qabbrev_tac ‘P = permutator d’ \\ + MATCH_MP_TAC (cj 2 subterm_subst_permutator_cong) \\ + qexistsl_tac [‘X’, ‘r’] >> simp []) + >> rpt GEN_TAC >> STRIP_TAC + >> qabbrev_tac ‘P = permutator d’ + >> ‘h::p IN ltree_paths (BT' X M r)’ by rw [ltree_paths_def] + >> Q.PAT_X_ASSUM ‘ltree_lookup (BT' X M r) (h::p) <> NONE’ MP_TAC + >> qabbrev_tac ‘Y = X UNION RANK r’ + >> reverse (Cases_on ‘solvable M’) + >- simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def] >> DISCH_TAC - >> Q.PAT_X_ASSUM ‘subterm X M (h::p) r <> NONE’ MP_TAC - >> Q_TAC (UNBETA_TAC [subterm_alt]) ‘subterm X M (h::p) r’ - >> STRIP_TAC - >> FIRST_X_ASSUM MATCH_MP_TAC >> art [] - >> CONJ_TAC - >- (MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) - >> reverse CONJ_TAC - >- (Q.PAT_X_ASSUM ‘v IN X UNION RANK r’ MP_TAC \\ - Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) - >> Know ‘LENGTH Ms = m’ - >- (simp [Once EQ_SYM_EQ, Abbr ‘m’] \\ - MATCH_MP_TAC hnf_children_size_alt \\ - qexistsl_tac [‘X’, ‘M’, ‘r’, ‘n’, ‘vs’, ‘M1’] >> simp []) + >> qabbrev_tac ‘M0 = principle_hnf M’ + >> qabbrev_tac ‘n = LAMl_size M0’ + >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ + >> ‘DISJOINT (set vs) (FV M)’ by METIS_TAC [subterm_disjoint_lemma] + >> ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ + >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, + “y :string”, “args :term list”)) ‘M1’ + >> ‘TAKE n vs = vs’ by rw [] + >> POP_ASSUM (rfs o wrap) + >> Q.PAT_X_ASSUM ‘DISJOINT (set vs) (FV M0)’ K_TAC + >> Know ‘~MEM v vs’ + >- (Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ + rw [Abbr ‘Y’, IN_UNION] + >- (Q.PAT_X_ASSUM ‘DISJOINT (set vs) X’ MP_TAC \\ + rw [DISJOINT_ALT']) \\ + Suff ‘DISJOINT (RANK r) (set vs)’ >- rw [DISJOINT_ALT] \\ + qunabbrev_tac ‘vs’ \\ + MATCH_MP_TAC DISJOINT_RANK_RNEWS' >> art []) >> DISCH_TAC - (* applying FV_subterm_lemma *) - >> Know ‘FV (EL h Ms) SUBSET FV M UNION set vs’ - >- (MATCH_MP_TAC FV_subterm_lemma \\ - qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘m’, ‘M1’] >> simp []) + >> Know ‘solvable ([P/v] M)’ + >- (MATCH_MP_TAC (cj 1 solvable_subst_permutator) \\ + qexistsl_tac [‘X’, ‘r’, ‘d’] >> simp [subterm_width_nil] \\ + MP_TAC (Q.SPECL [‘X’, ‘M’, ‘h::p’, ‘r’] subterm_width_first) \\ + simp [ltree_paths_def]) >> DISCH_TAC - >> STRIP_TAC - >> Know ‘v IN FV M UNION set vs’ >- METIS_TAC [SUBSET_DEF] - >> rw [IN_UNION] - >> Q.PAT_X_ASSUM ‘v IN X UNION RANK r’ MP_TAC - >> rw [IN_UNION] - >- (qunabbrev_tac ‘vs’ \\ - Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ \\ - Q.PAT_X_ASSUM ‘DISJOINT (set vs) X’ MP_TAC \\ - rw [DISJOINT_ALT']) - >> Suff ‘DISJOINT (RANK r) (set vs)’ >- rw [DISJOINT_ALT] - >> qunabbrev_tac ‘vs’ - >> MATCH_MP_TAC DISJOINT_RANK_RNEWS >> simp [] -QED - -Theorem subterm_once_fresh_tpm_cong[local] : - !p X M r v v'. FINITE X /\ FV M SUBSET X UNION RANK r /\ - subterm X M p r <> NONE /\ - v IN X UNION RANK r /\ - v' IN X UNION RANK r /\ v' # M - ==> subterm X (tpm [(v,v')] M) p r <> NONE /\ - subterm' X (tpm [(v,v')] M) p r = tpm [(v,v')] (subterm' X M p r) -Proof - rpt GEN_TAC >> STRIP_TAC - >> Know ‘v' # subterm' X M p r’ - >- (MATCH_MP_TAC FV_subterm_thm >> art []) + >> ‘M -h->* M0’ by METIS_TAC [principle_hnf_thm'] + >> Know ‘[P/v] M -h->* [P/v] M0’ >- PROVE_TAC [hreduce_substitutive] + >> ‘DISJOINT (set vs) (FV P)’ by rw [DISJOINT_ALT', FV_permutator, Abbr ‘P’] + >> simp [LAMl_SUB, appstar_SUB] + >> POP_ASSUM K_TAC (* DISJOINT (set vs) (FV P) *) + >> qabbrev_tac ‘args' = MAP [P/v] args’ + >> qabbrev_tac ‘m = LENGTH args’ + >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’, Abbr ‘m’] >> DISCH_TAC - >> simp [fresh_tpm_subst'] - >> MATCH_MP_TAC subterm_fresh_subst_cong >> art [] -QED - -Theorem subterm_fresh_tpm_cong_lemma[local] : - !pi X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - subterm X M p r <> NONE /\ - set (MAP FST pi) SUBSET X UNION RANK r /\ - set (MAP SND pi) SUBSET X UNION RANK r /\ - ALL_DISTINCT (MAP FST pi) /\ - ALL_DISTINCT (MAP SND pi) /\ - DISJOINT (set (MAP FST pi)) (set (MAP SND pi)) /\ - DISJOINT (set (MAP SND pi)) (FV M) - ==> subterm X (tpm pi M) p r <> NONE /\ - subterm' X (tpm pi M) p r = tpm pi (subterm' X M p r) -Proof - Induct_on ‘pi’ >- rw [] - >> simp [FORALL_PROD] - >> qx_genl_tac [‘u’, ‘v’] - >> rpt GEN_TAC - >> STRIP_TAC - >> ‘!t. tpm ((u,v)::pi) t = tpm [(u,v)] (tpm pi t)’ by rw [Once tpm_CONS] - >> POP_ORW - >> qabbrev_tac ‘N = tpm pi M’ - (* applying IH *) - >> Q.PAT_X_ASSUM ‘!X M p r. P’ (MP_TAC o Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’]) - >> simp [] - >> STRIP_TAC - >> POP_ASSUM (REWRITE_TAC o wrap o SYM) - >> MATCH_MP_TAC subterm_once_fresh_tpm_cong - >> simp [] - >> qabbrev_tac ‘vs = MAP FST pi’ - >> qabbrev_tac ‘vs' = MAP SND pi’ - >> ‘LENGTH vs' = LENGTH vs’ by rw [Abbr ‘vs’, Abbr ‘vs'’] - >> CONJ_TAC - >- (simp [Abbr ‘N’, SUBSET_DEF, FV_tpm] \\ - rpt STRIP_TAC \\ - qabbrev_tac ‘y = lswapstr (REVERSE pi) x’ \\ - ‘x = lswapstr pi y’ by rw [Abbr ‘y’] >> POP_ORW \\ - Know ‘pi = ZIP (vs,vs')’ - >- (simp [Abbr ‘vs’, Abbr ‘vs'’, LIST_EQ_REWRITE] \\ - rw [EL_ZIP, EL_MAP]) >> Rewr' \\ - Cases_on ‘MEM y vs’ - >- (Suff ‘lswapstr (ZIP (vs,vs')) y IN set vs'’ >- ASM_SET_TAC [] \\ - MATCH_MP_TAC MEM_lswapstr >> simp []) \\ - Cases_on ‘MEM y vs'’ - >- (Suff ‘lswapstr (ZIP (vs,vs')) y IN set vs’ >- ASM_SET_TAC [] \\ - MATCH_MP_TAC MEM_lswapstr' >> simp []) \\ - Suff ‘lswapstr (ZIP (vs,vs')) y = y’ >- ASM_SET_TAC [] \\ - MATCH_MP_TAC lswapstr_unchanged' \\ - simp [MAP_ZIP]) - >> simp [Abbr ‘N’] - >> Suff ‘lswapstr (REVERSE pi) v = v’ >- rw [] - >> MATCH_MP_TAC lswapstr_unchanged' - >> simp [MAP_REVERSE, MEM_REVERSE] -QED - -Theorem subterm_fresh_tpm_cong : - !pi X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - set (MAP FST pi) SUBSET X UNION RANK r /\ - set (MAP SND pi) SUBSET X UNION RANK r /\ - ALL_DISTINCT (MAP FST pi) /\ - ALL_DISTINCT (MAP SND pi) /\ - DISJOINT (set (MAP FST pi)) (set (MAP SND pi)) /\ - DISJOINT (set (MAP SND pi)) (FV M) - ==> (subterm X M p r = NONE <=> subterm X (tpm pi M) p r = NONE) /\ - (subterm X M p r <> NONE ==> - subterm' X (tpm pi M) p r = tpm pi (subterm' X M p r)) -Proof - reverse (rpt STRIP_TAC) - >- (MATCH_MP_TAC (cj 2 subterm_fresh_tpm_cong_lemma) >> art []) - >> reverse EQ_TAC - >- (CCONTR_TAC >> fs [] \\ - Q.PAT_X_ASSUM ‘subterm X (tpm pi M) p r = NONE’ MP_TAC \\ + >> Know ‘!i. i < m ==> FV (EL i args') SUBSET FV (EL i args)’ + >- (rw [Abbr ‘m’, Abbr ‘args'’, EL_MAP] \\ + qabbrev_tac ‘N = EL i args’ \\ + simp [FV_SUB] \\ + ‘FV P = {}’ by rw [Abbr ‘P’, FV_permutator] \\ simp [] \\ - MATCH_MP_TAC (cj 1 subterm_fresh_tpm_cong_lemma) >> art []) -(* stage work *) - >> CCONTR_TAC >> fs [] - >> qabbrev_tac ‘N = tpm pi M’ - >> Q.PAT_X_ASSUM ‘subterm X M p r = NONE’ MP_TAC - >> ‘M = tpm (REVERSE pi) N’ by rw [Abbr ‘N’] - >> simp [] - >> qabbrev_tac ‘xs = MAP FST pi’ - >> qabbrev_tac ‘ys = MAP SND pi’ - >> ‘LENGTH ys = LENGTH xs’ by rw [Abbr ‘xs’, Abbr ‘ys’] - >> Know ‘pi = ZIP (xs,ys)’ - >- (rw [Abbr ‘xs’, Abbr ‘ys’, ZIP_MAP] \\ - rw [LIST_EQ_REWRITE, EL_MAP]) + Cases_on ‘v IN FV N’ >> SET_TAC []) >> DISCH_TAC - >> simp [REVERSE_ZIP] - >> qabbrev_tac ‘xs' = REVERSE xs’ - >> qabbrev_tac ‘ys' = REVERSE ys’ - >> ‘LENGTH ys' = LENGTH xs'’ by rw [Abbr ‘xs'’, Abbr ‘ys'’] - (* applying pmact_flip_args_all *) - >> ‘tpm (ZIP (xs',ys')) N = tpm (ZIP (ys',xs')) N’ by rw [pmact_flip_args_all] - >> POP_ORW - >> qabbrev_tac ‘pi' = ZIP (ys',xs')’ - >> MATCH_MP_TAC (cj 1 subterm_fresh_tpm_cong_lemma) - >> simp [Abbr ‘pi'’, MAP_ZIP] - (* applying FV_tpm_lemma *) - >> CONJ_TAC - >- (qunabbrev_tac ‘N’ \\ - MATCH_MP_TAC FV_tpm_lemma \\ - Q.EXISTS_TAC ‘r’ >> simp [MAP_ZIP] \\ - ASM_SET_TAC []) - >> CONJ_TAC - >- (Q.PAT_X_ASSUM ‘set ys SUBSET X UNION RANK r’ MP_TAC \\ - rw [SUBSET_DEF, Abbr ‘ys'’]) - >> CONJ_TAC - >- (Q.PAT_X_ASSUM ‘set xs SUBSET X UNION RANK r’ MP_TAC \\ - rw [SUBSET_DEF, Abbr ‘xs'’]) - >> CONJ_TAC >- rw [ALL_DISTINCT_REVERSE, Abbr ‘ys'’] - >> CONJ_TAC >- rw [ALL_DISTINCT_REVERSE, Abbr ‘xs'’] - >> CONJ_TAC - >- (rw [DISJOINT_ALT', Abbr ‘xs'’] \\ - Suff ‘DISJOINT (set xs) (set ys')’ >- rw [DISJOINT_ALT] \\ - rw [DISJOINT_ALT', Abbr ‘ys'’] \\ - Q.PAT_X_ASSUM ‘DISJOINT (set xs) (set ys)’ MP_TAC \\ - rw [DISJOINT_ALT']) + >> Q.PAT_X_ASSUM ‘ltree_lookup (BT' X M r) (h::p) <> NONE’ + (MP_TAC o REWRITE_RULE [BT_def, BT_generator_def, ltree_unfold]) + >> simp [principle_hnf_beta_reduce, hnf_appstar, LMAP_fromList, + ltree_lookup_def, LNTH_fromList] + >> Cases_on ‘h < m’ >> simp [EL_MAP, GSYM BT_def] + >> DISCH_TAC + >> Know ‘subterm_width M (h::p) <= d <=> + m <= d /\ subterm_width (EL h args) p <= d’ + >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ + qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘vs’, ‘M1’] >> simp []) + >> DISCH_THEN (rfs o wrap) + >> qabbrev_tac ‘N = EL h args’ (* stage work *) - >> Suff ‘DISJOINT (set xs) (FV N)’ - >- rw [DISJOINT_ALT, Abbr ‘xs'’] - >> POP_ASSUM K_TAC - >> qunabbrevl_tac [‘xs'’, ‘ys'’] - >> simp [Abbr ‘N’] - >> POP_ASSUM K_TAC (* pi = ZIP (xs,ys) *) - >> ONCE_REWRITE_TAC [DISJOINT_SYM] - >> Know ‘tpm (ZIP (xs,ys)) M = M ISUB ZIP (MAP VAR ys,xs)’ - >- (MATCH_MP_TAC fresh_tpm_isub' >> art []) - >> Rewr' - >> MATCH_MP_TAC FV_renaming_disjoint >> art [] -QED - -(* Lemma 10.3.6 (ii) [1, p.247]: - - NOTE: The construction of ‘pi’ needs a fixed ltree path ‘p’, so that we can - collect the maximum number of children in all nodes along ‘p’. In another - word, there exists no universal ‘pi’ for which the conclusion holds for - arbitrary ‘p’. - - NOTE: Added ‘subterm X M p <> NONE’ to antecedents so that ‘subterm' X M p’ - is defined/specified. ‘subterm X (apply pi M) p <> NONE’ can be derived. - - NOTE: ‘p <> []’ must be added into antecedents, otherwise the statement - becomes: - - [...] |- !X M. ?pi. Boehm_transform pi /\ is_ready (apply pi M) /\ - ?fm. apply pi M == fm ' M - - which is impossible if M is not already "is_ready". + >> reverse (Cases_on ‘y = v’) + >- (Q.PAT_X_ASSUM ‘[P/v] M -h->* _’ MP_TAC \\ + simp [] >> DISCH_TAC \\ + qabbrev_tac ‘M0' = LAMl vs (VAR y @* args')’ \\ + ‘hnf M0'’ by rw [hnf_appstar, Abbr ‘M0'’] \\ + ‘principle_hnf ([P/v] M) = M0'’ by METIS_TAC [principle_hnf_thm'] \\ + Q.PAT_X_ASSUM ‘hnf M0'’ K_TAC \\ + Q.PAT_X_ASSUM ‘M0 = LAMl vs (VAR y @* args)’ (ASSUME_TAC o SYM) \\ + Q.PAT_X_ASSUM ‘M1 = VAR y @* args’ (ASSUME_TAC o SYM) \\ + ‘hnf_children M1 = args’ by rw [hnf_children_hnf] \\ + ‘LAMl_size M0' = n’ by rw [Abbr ‘M0'’, LAMl_size_hnf] \\ + ‘principle_hnf (M0' @* MAP VAR vs) = VAR y @* args'’ + by rw [Abbr ‘M0'’, principle_hnf_beta_reduce, hnf_appstar] \\ + STRONG_CONJ_TAC + >- (simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def, + LNTH_fromList, EL_MAP] \\ + simp [GSYM BT_def, EL_MAP, Abbr ‘args'’] \\ + FIRST_X_ASSUM (MATCH_MP_TAC o cj 1) >> art [] \\ + CONJ_TAC + >- (qunabbrev_tac ‘N’ \\ + MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] \\ + Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ + simp []) \\ + Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ + qunabbrev_tac ‘Y’ \\ + Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ + >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [RANK_MONO]) >> DISCH_TAC \\ + Know ‘subterm_width ([P/v] M) (h::p) <= d <=> + m <= d /\ subterm_width (EL h args') p <= d’ + >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ + qexistsl_tac [‘X’, ‘r’, ‘M0'’, ‘n’, ‘vs’, ‘VAR y @* args'’] \\ + simp [principle_hnf_beta_reduce, ltree_paths_def] \\ + CONJ_TAC + >- (rw [FV_SUB, Abbr ‘P’, FV_permutator] \\ + MATCH_MP_TAC SUBSET_TRANS >> Q.EXISTS_TAC ‘FV M’ >> art [] \\ + SET_TAC []) \\ + fs [Abbr ‘m’, Abbr ‘args'’, Abbr ‘M0'’]) >> Rewr' \\ + simp [Abbr ‘args'’, EL_MAP] \\ + FIRST_X_ASSUM (MATCH_MP_TAC o cj 2) >> art [] \\ + Q.EXISTS_TAC ‘SUC r’ >> art [] \\ + CONJ_TAC + >- (qunabbrev_tac ‘N’ \\ + MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] \\ + Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ + simp []) \\ + Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ + qunabbrev_tac ‘Y’ \\ + Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ + >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [RANK_MONO]) + (* stage work *) + >> POP_ASSUM (rfs o wrap o SYM) (* [P/y] from now on *) + >> Q.PAT_X_ASSUM ‘[P/y] M -h->* _’ MP_TAC + >> simp [Abbr ‘P’] + >> DISCH_TAC (* [permutator d/v] M -h->* ... *) + (* NOTE: New frash variables xs and y will be asserted here: - NOTE: Added ‘p IN ltree_paths (BT X (apply pi M))’ into conclusions for the - needs in the user theorem. + P @* args' -h->* LAMl xs (LAM z (VAR z @* args' @* MAP VAR xs)), - NOTE: Extended the conclusion with ‘!q. q <<= p /\ q <> []’ + thus ‘principle_hnf ([P/y] M) = LAMl (vs ++ xs ++ [z]) (VAR z @* ...)’ - NOTE: ‘FV (apply pi M) SUBSET X UNION RANK (SUC r)’ is added into the - conclusions for the needs of Boehm_out_lemma. - *) -Theorem Boehm_transform_exists_lemma : - !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - p <> [] /\ subterm X M p r <> NONE ==> - ?pi. Boehm_transform pi /\ is_ready' (apply pi M) /\ - FV (apply pi M) SUBSET X UNION RANK (SUC r) /\ - ?v P. closed P /\ - !q. q <<= p /\ q <> [] ==> - subterm X (apply pi M) q r <> NONE /\ - subterm' X (apply pi M) q r = [P/v] (subterm' X M q r) -Proof - rpt STRIP_TAC - >> ‘p IN ltree_paths (BT' X M r)’ by PROVE_TAC [subterm_imp_ltree_paths] - >> ‘(!q. q <<= p ==> subterm X M q r <> NONE) /\ - !q. q <<= FRONT p ==> solvable (subterm' X M q r)’ - by (MATCH_MP_TAC subterm_solvable_lemma >> art []) - >> Know ‘solvable M’ - >- (POP_ASSUM (MP_TAC o Q.SPEC ‘[]’) >> rw []) - >> DISCH_TAC - (* M0 is now meaningful since M is solvable *) - >> qabbrev_tac ‘M0 = principle_hnf M’ - >> ‘hnf M0’ by PROVE_TAC [hnf_principle_hnf'] - >> qabbrev_tac ‘n = LAMl_size M0’ - (* NOTE: here the excluded list must contain ‘FV M’. Just ‘FV M0’ doesn't - work later, when calling the important [principle_hnf_denude_thm]. - This is conflicting with BT_generator_def and subterm_def. + Here LENGTH xs = d - m. Let n' be the LAMl_size of the above hnf. *) - >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ - >> ‘DISJOINT (set vs) (FV M)’ by METIS_TAC [subterm_disjoint_lemma] - >> Know ‘?y args. M0 = LAMl (TAKE n vs) (VAR y @* args)’ - >- (qunabbrev_tac ‘n’ \\ - ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] \\ - irule (iffLR hnf_cases_shared) >> rw [] \\ - MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘FV M’ >> art [] \\ - qunabbrev_tac ‘M0’ >> MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) + >> qabbrev_tac ‘n' = n + (d - m) + 1’ + >> Q_TAC (RNEWS_TAC (“zs :string list”, “r :num”, “n' :num”)) ‘X’ + >> MP_TAC (Q.SPECL [‘set zs’, ‘d’, ‘args'’] hreduce_permutator_thm) + >> impl_tac >- rpt (rw []) + >> REWRITE_TAC [DISJOINT_UNION] >> STRIP_TAC - (* eliminate ‘TAKE n vs’ *) - >> ‘TAKE n vs = vs’ by rw [] - >> POP_ASSUM (REV_FULL_SIMP_TAC std_ss o wrap) - >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ - >> Know ‘M1 = VAR y @* args’ - >- (qunabbrev_tac ‘M1’ >> POP_ORW \\ - MATCH_MP_TAC principle_hnf_beta_reduce >> rw [hnf_appstar]) - >> DISCH_TAC - >> qabbrev_tac ‘m = LENGTH args’ - (* using ‘subterm_width’ and applying subterm_width_thm *) - >> qabbrev_tac ‘d = subterm_width M p’ - >> Know ‘m <= d’ - >- (MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_first) \\ - rw [Abbr ‘d’]) - >> DISCH_TAC - (* p1 is the first Boehm transformation for removing abstractions of M0 *) - >> qabbrev_tac ‘p1 = MAP rightctxt (REVERSE (MAP VAR vs))’ - >> ‘Boehm_transform p1’ by rw [Abbr ‘p1’, MAP_MAP_o, GSYM MAP_REVERSE] - >> ‘apply p1 M0 == M1’ by rw [Abbr ‘p1’, Boehm_apply_MAP_rightctxt'] - (* stage work (all correct until here) - - Now we define the permutator P (and then p2). This requires q + 1 fresh - variables. The excluded list is at least X and FV M, and then ‘vs’. - But since P is a closed term, these fresh variables seem irrelevant... - *) + >> rename1 ‘ALL_DISTINCT (SNOC z xs)’ (* y' -> z *) >> qabbrev_tac ‘P = permutator d’ - >> ‘closed P’ by rw [Abbr ‘P’, closed_permutator] - >> qabbrev_tac ‘p2 = [[P/y]]’ - >> ‘Boehm_transform p2’ by rw [Abbr ‘p2’] - >> ‘apply p2 M1 = P @* MAP [P/y] args’ by rw [Abbr ‘p2’, appstar_SUB] - >> qabbrev_tac ‘args' = MAP [P/y] args’ - >> Know ‘!i. i < m ==> FV (EL i args') SUBSET FV (EL i args)’ - >- (rw [Abbr ‘args'’, EL_MAP, FV_SUB] \\ - fs [closed_def]) + >> Know ‘LAMl vs (P @* args') -h->* + LAMl vs (LAMl xs (LAM z (VAR z @* args' @* MAP VAR xs)))’ + >- (rw [hreduce_LAMl]) + >> qmatch_abbrev_tac ‘LAMl vs (P @* args') -h->* t ==> _’ >> DISCH_TAC - >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’, Abbr ‘m’] - (* NOTE: Z contains ‘vs’ in addition to X and FV M *) - >> qabbrev_tac ‘Z = X UNION FV M UNION set vs’ - >> ‘FINITE Z’ by (rw [Abbr ‘Z’] >> rw []) - >> Know ‘solvable (M0 @* MAP VAR vs)’ - >- (MATCH_MP_TAC solvable_appstar' \\ - qexistsl_tac [‘X’, ‘M’, ‘r’, ‘n’] >> simp []) + >> ‘hnf t’ by rw [Abbr ‘t’, hnf_appstar, hnf_LAMl] + >> ‘[P/y] M -h->* t’ by PROVE_TAC [hreduce_TRANS] + >> ‘principle_hnf ([P/y] M) = t’ by METIS_TAC [principle_hnf_thm'] + (* cleanup *) + >> Q.PAT_X_ASSUM ‘P @* args' -h->* _’ K_TAC + >> Q.PAT_X_ASSUM ‘LAMl vs _ -h->* t’ K_TAC + >> Q.PAT_X_ASSUM ‘[P/y] M -h->* LAMl vs _’ K_TAC + >> Q.PAT_X_ASSUM ‘[P/y] M -h->* t’ K_TAC + >> Q.PAT_X_ASSUM ‘hnf t’ K_TAC + >> qunabbrev_tac ‘t’ + (* stage work *) + >> POP_ASSUM MP_TAC + >> ‘!t. LAMl vs (LAMl xs (LAM z t)) = LAMl (vs ++ (SNOC z xs)) t’ + by rw [LAMl_APPEND] + >> POP_ORW + >> REWRITE_TAC [GSYM appstar_APPEND] + >> qabbrev_tac ‘args'' = args' ++ MAP VAR xs’ + >> ‘LENGTH args'' = d’ by rw [Abbr ‘args''’] + >> qabbrev_tac ‘xs' = SNOC z xs’ + >> ‘LENGTH xs' = d - m + 1’ by rw [Abbr ‘xs'’] + >> qabbrev_tac ‘vs' = vs ++ xs'’ + >> DISCH_TAC (* principle_hnf ([P/y] M) = ... *) + >> Know ‘LENGTH vs' = n'’ + >- (qunabbrevl_tac [‘n’, ‘n'’, ‘vs'’, ‘xs'’] \\ + Q.PAT_X_ASSUM ‘M0 = _’ (REWRITE_TAC o wrap) \\ + simp [LAMl_size_LAMl]) >> DISCH_TAC - >> Know ‘FV M1 SUBSET Z’ - >- (MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘FV M0 UNION set vs’ \\ - reverse CONJ_TAC - >- (qunabbrev_tac ‘Z’ \\ - Suff ‘FV M0 SUBSET FV M’ >- SET_TAC [] \\ - qunabbrev_tac ‘M0’ \\ - MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) \\ - ‘FV M0 UNION set vs = FV (M0 @* MAP VAR vs)’ by rw [FV_appstar_MAP_VAR] \\ - POP_ORW \\ - qunabbrev_tac ‘M1’ \\ - MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) - >> DISCH_TAC - >> qabbrev_tac ‘z = SUC (string_width Z)’ - >> qabbrev_tac ‘l = alloc r z (z + d - m + 1)’ - >> Know ‘ALL_DISTINCT l /\ LENGTH l = d - m + 1’ - >- rw [Abbr ‘l’, alloc_thm] - >> STRIP_TAC - >> Know ‘DISJOINT (set l) Z’ - >- (rw [Abbr ‘l’, Abbr ‘z’, DISJOINT_ALT', alloc_def, MEM_MAP] \\ - ONCE_REWRITE_TAC [TAUT ‘~P \/ Q \/ ~R <=> P /\ R ==> Q’] \\ - STRIP_TAC \\ - ‘FINITE Z’ by rw [Abbr ‘Z’] \\ - MP_TAC (Q.SPECL [‘x’, ‘Z’] string_width_thm) >> rw []) - >> DISCH_TAC - (* now recover the old definition of Y *) - >> Know ‘DISJOINT (set l) (FV M1)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘Z’ >> art []) - >> ASM_REWRITE_TAC [FV_appstar, FV_thm] - >> DISCH_THEN (STRIP_ASSUME_TAC o (REWRITE_RULE [DISJOINT_UNION'])) - >> Q.PAT_X_ASSUM ‘DISJOINT (set l) {y}’ (* ~MEM y l *) - (STRIP_ASSUME_TAC o (SIMP_RULE (srw_ss()) [DISJOINT_ALT'])) - >> ‘l <> []’ by rw [NOT_NIL_EQ_LENGTH_NOT_0] - >> qabbrev_tac ‘as = FRONT l’ - >> ‘LENGTH as = d - m’ by rw [Abbr ‘as’, LENGTH_FRONT] - >> qabbrev_tac ‘b = LAST l’ - >> Know ‘l = SNOC b as’ - >- (ASM_SIMP_TAC std_ss [Abbr ‘as’, Abbr ‘b’, SNOC_LAST_FRONT]) + (* applying NEWS_prefix *) + >> Know ‘vs <<= zs’ + >- (qunabbrevl_tac [‘vs’, ‘zs’] \\ + MATCH_MP_TAC RNEWS_prefix >> rw [Abbr ‘n'’]) + >> DISCH_THEN (STRIP_ASSUME_TAC o (REWRITE_RULE [IS_PREFIX_APPEND])) + >> rename1 ‘zs = vs ++ ys'’ (* l -> ys' *) + >> Know ‘LENGTH ys' = d - m + 1’ + >- (Know ‘LENGTH zs = LENGTH (vs ++ ys')’ >- POP_ASSUM (REWRITE_TAC o wrap) \\ + simp [Abbr ‘n'’]) >> DISCH_TAC - >> qabbrev_tac ‘p3 = MAP rightctxt (REVERSE (MAP VAR l))’ - >> ‘Boehm_transform p3’ by rw [Abbr ‘p3’, MAP_MAP_o, GSYM MAP_REVERSE] - (* applying permutator_thm *) - >> Know ‘apply p3 (P @* args') == VAR b @* args' @* MAP VAR as’ - >- (simp [Abbr ‘p3’, Abbr ‘P’, rightctxt_thm, MAP_SNOC, - Boehm_apply_MAP_rightctxt'] \\ - REWRITE_TAC [GSYM appstar_APPEND, APPEND_ASSOC] \\ - MATCH_MP_TAC permutator_thm >> rw []) + >> Know ‘!N. MEM N args' ==> DISJOINT (FV N) (set ys')’ + >- (Q.X_GEN_TAC ‘x’ \\ + simp [MEM_EL] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘i’ STRIP_ASSUME_TAC) \\ + POP_ORW \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘FV (EL i args)’ >> simp [] \\ + (* applying FV_subterm_lemma *) + Know ‘FV (EL i args) SUBSET FV M UNION set vs’ + >- (MATCH_MP_TAC FV_subterm_lemma \\ + qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘m’, ‘M1’] >> simp []) \\ + DISCH_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘FV M UNION set vs’ >> art [] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘X UNION RANK r UNION set vs’ \\ + reverse CONJ_TAC >- ASM_SET_TAC [] \\ + simp [DISJOINT_ALT'] \\ + rpt CONJ_TAC + >- (NTAC 2 STRIP_TAC \\ + Q.PAT_X_ASSUM ‘DISJOINT (set zs) X’ MP_TAC \\ + rw [DISJOINT_ALT]) + >- (NTAC 2 STRIP_TAC \\ + MP_TAC (Q.SPECL [‘r’, ‘n'’, ‘X’] DISJOINT_RANK_RNEWS') \\ + rw [DISJOINT_ALT', IN_UNION]) \\ + Q.PAT_X_ASSUM ‘ALL_DISTINCT zs’ MP_TAC \\ + rw [ALL_DISTINCT_APPEND] >> METIS_TAC []) >> DISCH_TAC - (* pre-final stage *) - >> Q.EXISTS_TAC ‘p3 ++ p2 ++ p1’ - >> CONJ_ASM1_TAC - >- (MATCH_MP_TAC Boehm_transform_APPEND >> art [] \\ - MATCH_MP_TAC Boehm_transform_APPEND >> art []) - >> Know ‘apply (p3 ++ p2 ++ p1) M == VAR b @* args' @* MAP VAR as’ - >- (MATCH_MP_TAC lameq_TRANS \\ - Q.EXISTS_TAC ‘apply (p3 ++ p2 ++ p1) M0’ \\ - CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong \\ - CONJ_TAC >- art [] \\ - qunabbrev_tac ‘M0’ \\ - MATCH_MP_TAC lameq_SYM \\ - MATCH_MP_TAC lameq_principle_hnf' >> art []) \\ - ONCE_REWRITE_TAC [Boehm_apply_APPEND] \\ - MATCH_MP_TAC lameq_TRANS \\ - Q.EXISTS_TAC ‘apply (p3 ++ p2) M1’ \\ - CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong >> art [] \\ - MATCH_MP_TAC Boehm_transform_APPEND >> art []) \\ - ONCE_REWRITE_TAC [Boehm_apply_APPEND] \\ - MATCH_MP_TAC lameq_TRANS \\ - Q.EXISTS_TAC ‘apply p3 (P @* args')’ >> art [] \\ - MATCH_MP_TAC Boehm_apply_lameq_cong >> rw []) + >> qabbrev_tac ‘t = VAR z @* args''’ + >> ‘hnf t’ by rw [Abbr ‘t’, hnf_appstar] + >> qabbrev_tac ‘M0' = LAMl vs' t’ + >> ‘LAMl_size M0' = n'’ by rw [Abbr ‘M0'’, Abbr ‘t’] + >> qabbrev_tac ‘M1' = principle_hnf (M0' @* MAP VAR zs)’ + >> Know ‘M1' = tpm (ZIP (xs',ys')) t’ + >- (simp [Abbr ‘M0'’, Abbr ‘M1'’, Abbr ‘vs'’, GSYM APPEND_ASSOC, appstar_APPEND, + LAMl_APPEND] \\ + Know ‘principle_hnf (LAMl vs (LAMl xs' t) @* MAP VAR vs @* MAP VAR ys') = + principle_hnf (LAMl xs' t @* MAP VAR ys')’ + >- (MATCH_MP_TAC principle_hnf_hreduce \\ + simp [hreduce_BETA_extended]) >> Rewr' \\ + MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ + CONJ_TAC >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT zs’ MP_TAC \\ + Q.PAT_X_ASSUM ‘zs = vs ++ ys'’ (ONCE_REWRITE_TAC o wrap) \\ + simp [ALL_DISTINCT_APPEND]) \\ + CONJ_ASM1_TAC + >- (ONCE_REWRITE_TAC [DISJOINT_SYM] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set zs’ >> art [] \\ + rw [LIST_TO_SET_APPEND]) \\ + simp [Abbr ‘t’, Abbr ‘args''’, FV_appstar] \\ + CONJ_TAC (* ~MEM z ys' *) + >- (POP_ASSUM MP_TAC \\ + rw [DISJOINT_ALT, Abbr ‘xs'’]) \\ + reverse CONJ_TAC (* MAP VAR xs *) + >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs') (set ys')’ MP_TAC \\ + rw [Abbr ‘xs'’, DISJOINT_ALT]) \\ + Q.X_GEN_TAC ‘s’ >> simp [MEM_MAP] \\ + rpt STRIP_TAC \\ + Q.PAT_X_ASSUM ‘s = FV x’ (REWRITE_TAC o wrap) \\ + FIRST_X_ASSUM MATCH_MP_TAC >> art []) >> DISCH_TAC + >> Q.PAT_X_ASSUM ‘zs = vs ++ ys'’ (ASSUME_TAC o SYM) (* stage work *) - >> Know ‘principle_hnf (apply (p3 ++ p2 ++ p1) M) = VAR b @* args' @* MAP VAR as’ - >- (Q.PAT_X_ASSUM ‘Boehm_transform (p3 ++ p2 ++ p1)’ K_TAC \\ - Q.PAT_X_ASSUM ‘Boehm_transform p1’ K_TAC \\ - Q.PAT_X_ASSUM ‘Boehm_transform p2’ K_TAC \\ - Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC \\ - Q.PAT_X_ASSUM ‘apply p1 M0 == M1’ K_TAC \\ - Q.PAT_X_ASSUM ‘apply p2 M1 = P @* args'’ K_TAC \\ - Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC \\ - (* preparing for principle_hnf_denude_thm *) - Q.PAT_X_ASSUM ‘apply (p3 ++ p2 ++ p1) M == _’ MP_TAC \\ - simp [Boehm_apply_APPEND, Abbr ‘p1’, Abbr ‘p2’, Abbr ‘p3’, - Boehm_apply_MAP_rightctxt'] \\ - Know ‘[P/y] (M @* MAP VAR vs) @* MAP VAR (SNOC b as) = - [P/y] (M @* MAP VAR vs @* MAP VAR (SNOC b as))’ - >- (simp [appstar_SUB] \\ - Suff ‘MAP [P/y] (MAP VAR (SNOC b as)) = MAP VAR (SNOC b as)’ >- Rewr \\ - Q.PAT_X_ASSUM ‘l = SNOC b as’ (ONCE_REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘LENGTH l = d - m + 1’ K_TAC \\ - rw [LIST_EQ_REWRITE, EL_MAP] \\ - MATCH_MP_TAC lemma14b \\ - REWRITE_TAC [FV_thm, IN_SING] \\ - Q.PAT_X_ASSUM ‘~MEM y l’ MP_TAC \\ - rw [MEM_EL] >> METIS_TAC []) >> Rewr' \\ - DISCH_TAC (* [P/y] ... == ... *) \\ - (* applying principle_hnf_permutator *) - Know ‘VAR b @* args' @* MAP VAR as = - principle_hnf ([P/y] (VAR y @* args @* MAP VAR (SNOC b as)))’ - >- (ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ - simp [appstar_SUB, appstar_SNOC, MAP_SNOC] \\ - Know ‘MAP [P/y] (MAP VAR as) = MAP VAR as’ - >- (Q.PAT_X_ASSUM ‘LENGTH as = _’ K_TAC \\ - rw [LIST_EQ_REWRITE, EL_MAP] \\ - MATCH_MP_TAC lemma14b >> rw [] \\ - Q.PAT_X_ASSUM ‘~MEM y (SNOC b as)’ MP_TAC \\ - rw [MEM_EL] >> PROVE_TAC []) >> Rewr' \\ - Know ‘[P/y] (VAR b) = VAR b’ - >- (MATCH_MP_TAC lemma14b >> fs [MEM_SNOC, IN_UNION]) >> Rewr' \\ - simp [Abbr ‘P’, GSYM appstar_APPEND] \\ - MATCH_MP_TAC principle_hnf_permutator >> rw []) >> Rewr' \\ - (* applying principle_hnf_SUB_cong *) - MATCH_MP_TAC principle_hnf_SUB_cong \\ - CONJ_TAC (* has_hnf #1 *) - >- (simp [GSYM solvable_iff_has_hnf, GSYM appstar_APPEND] \\ - ‘M0 == M’ by rw [lameq_principle_hnf', Abbr ‘M0’] \\ - ‘M0 @* (MAP VAR vs ++ MAP VAR (SNOC b as)) == - M @* (MAP VAR vs ++ MAP VAR (SNOC b as))’ by rw [lameq_appstar_cong] \\ - Suff ‘solvable (M0 @* (MAP VAR vs ++ MAP VAR (SNOC b as)))’ - >- PROVE_TAC [lameq_solvable_cong] \\ - NTAC 2 (POP_ASSUM K_TAC) \\ - REWRITE_TAC [appstar_APPEND] \\ - qabbrev_tac ‘M0' = M0 @* MAP VAR vs’ \\ - ‘M0' == M1’ by rw [Abbr ‘M0'’] \\ - ‘M0' @* MAP VAR (SNOC b as) == M1 @* MAP VAR (SNOC b as)’ - by rw [lameq_appstar_cong] \\ - Suff ‘solvable (M1 @* MAP VAR (SNOC b as))’ - >- PROVE_TAC [lameq_solvable_cong] \\ - REWRITE_TAC [solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ - CONJ_TAC (* has_hnf #2 *) - >- (REWRITE_TAC [GSYM solvable_iff_has_hnf] \\ - Suff ‘solvable (VAR b @* args' @* MAP VAR as)’ - >- PROVE_TAC [lameq_solvable_cong] \\ - REWRITE_TAC [solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ - CONJ_TAC (* has_hnf # 3 *) - >- (simp [appstar_SUB, MAP_SNOC] \\ - Know ‘MAP [P/y] (MAP VAR as) = MAP VAR as’ - >- (Q.PAT_X_ASSUM ‘LENGTH as = _’ K_TAC \\ - rw [LIST_EQ_REWRITE, EL_MAP] \\ - MATCH_MP_TAC lemma14b >> rw [] \\ - Q.PAT_X_ASSUM ‘~MEM y (SNOC b as)’ MP_TAC \\ - rw [MEM_EL] >> PROVE_TAC []) >> Rewr' \\ - Know ‘[P/y] (VAR b) = VAR b’ - >- (MATCH_MP_TAC lemma14b >> fs [MEM_SNOC, IN_UNION]) >> Rewr' \\ - simp [Abbr ‘P’, GSYM appstar_APPEND] \\ - REWRITE_TAC [GSYM solvable_iff_has_hnf] \\ - Know ‘permutator d @* (args' ++ MAP VAR as) @@ VAR b == - VAR b @* (args' ++ MAP VAR as)’ - >- (MATCH_MP_TAC permutator_thm >> rw []) >> DISCH_TAC \\ - Suff ‘solvable (VAR b @* (args' ++ MAP VAR as))’ - >- PROVE_TAC [lameq_solvable_cong] \\ - REWRITE_TAC [solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ - (* applying the celebrating principle_hnf_denude_thm - - NOTE: here ‘DISJOINT (set vs) (FV M)’ is required, and this means that - ‘vs’ must exclude ‘FV M’ instead of just ‘FV M0’. - *) - MATCH_MP_TAC principle_hnf_denude_thm >> rw []) - >> DISCH_TAC - >> simp [is_ready', GSYM CONJ_ASSOC] - (* extra subgoal: solvable (apply (p3 ++ p2 ++ p1) M) *) - >> ONCE_REWRITE_TAC [TAUT ‘P /\ Q /\ R <=> Q /\ P /\ R’] - >> CONJ_ASM1_TAC - >- (Suff ‘solvable (VAR b @* args' @* MAP VAR as)’ - >- PROVE_TAC [lameq_solvable_cong] \\ - REWRITE_TAC [solvable_iff_has_hnf] \\ - MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) - (* applying is_ready_alt *) - >> CONJ_TAC - >- (simp [is_ready_alt] \\ - qexistsl_tac [‘b’, ‘args' ++ MAP VAR as’] \\ - CONJ_TAC - >- (MP_TAC (Q.SPEC ‘VAR b @* args' @* MAP VAR as’ - (MATCH_MP principle_hnf_thm' - (ASSUME “solvable (apply (p3 ++ p2 ++ p1) M)”))) \\ - simp [appstar_APPEND]) \\ - simp [] (* now two EVERY *) \\ - reverse CONJ_TAC - >- (rw [EVERY_MEM, Abbr ‘b’, Abbr ‘as’, MEM_MAP] >> rw [FV_thm] \\ - Q.PAT_X_ASSUM ‘ALL_DISTINCT l’ MP_TAC \\ - Q.PAT_X_ASSUM ‘l = SNOC (LAST l) (FRONT l)’ (ONCE_REWRITE_TAC o wrap) \\ - rw [ALL_DISTINCT_SNOC] >> PROVE_TAC []) \\ - rw [EVERY_MEM, MEM_MAP] \\ - qabbrev_tac ‘Y = BIGUNION (IMAGE FV (set args))’ \\ - rfs [LIST_TO_SET_SNOC] \\ - Suff ‘FV e SUBSET Y’ >- METIS_TAC [SUBSET_DEF] \\ + >> STRONG_CONJ_TAC + >- (simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def, + LNTH_fromList, EL_MAP] \\ + REWRITE_TAC [GSYM BT_def] \\ + simp [hnf_children_tpm, EL_MAP] \\ + simp [Abbr ‘t’, hnf_children_hnf] \\ + Know ‘h < LENGTH args''’ + >- (Q_TAC (TRANS_TAC LESS_LESS_EQ_TRANS) ‘m’ >> art [] \\ + rw [Abbr ‘args''’]) >> Rewr \\ + ‘EL h args'' = EL h args'’ by rw [Abbr ‘args''’, EL_APPEND1] \\ + POP_ORW \\ + Know ‘tpm (ZIP (xs',ys')) (EL h args') = EL h args'’ + >- (MATCH_MP_TAC tpm_unchanged >> simp [MAP_ZIP] \\ + ONCE_REWRITE_TAC [DISJOINT_SYM] \\ + CONJ_TAC \\ (* 2 subgoals, same tactics *) + FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) >> Rewr' \\ + simp [Abbr ‘args'’, EL_MAP] \\ + FIRST_X_ASSUM (MATCH_MP_TAC o cj 1) >> art [] \\ qunabbrev_tac ‘Y’ \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘BIGUNION (IMAGE FV (set args'))’ \\ - reverse CONJ_TAC - >- (rw [SUBSET_DEF, IN_BIGUNION_IMAGE, MEM_EL] \\ - Q.EXISTS_TAC ‘EL n args’ \\ - CONJ_TAC >- (Q.EXISTS_TAC ‘n’ >> art []) \\ - POP_ASSUM MP_TAC \\ - Suff ‘FV (EL n args') SUBSET FV (EL n args)’ >- METIS_TAC [SUBSET_DEF] \\ - FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ - rw [SUBSET_DEF, IN_BIGUNION_IMAGE] \\ - Q.EXISTS_TAC ‘e’ >> art []) - (* extra goal *) - >> CONJ_TAC - >- (Q.PAT_X_ASSUM ‘apply (p3 ++ p2 ++ p1) M == _’ K_TAC \\ - Q.PAT_X_ASSUM ‘principle_hnf (apply (p3 ++ p2 ++ p1) M) = _’ K_TAC \\ - Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC \\ - rpt (Q.PAT_X_ASSUM ‘Boehm_transform _’ K_TAC) \\ - Q.PAT_X_ASSUM ‘solvable (apply (p3 ++ p2 ++ p1) M)’ K_TAC \\ - simp [Boehm_apply_APPEND, Abbr ‘p1’, Abbr ‘p2’, Abbr ‘p3’, - Boehm_apply_MAP_rightctxt'] \\ - POP_ASSUM (ONCE_REWRITE_TAC o wrap o SYM) \\ reverse CONJ_TAC - >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW r’ \\ - rw [Abbr ‘l’, alloc_SUBSET_ROW] \\ - Suff ‘ROW r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [ROW_SUBSET_RANK]) \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘FV (M @* MAP VAR vs)’ \\ - CONJ_TAC >- (MATCH_MP_TAC FV_SUB_SUBSET >> art []) \\ - simp [FV_appstar] \\ - reverse CONJ_TAC - >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW r’ \\ - rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ - Suff ‘ROW r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [ROW_SUBSET_RANK]) \\ - Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ >> art [] \\ - Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) - (* stage work, there's the textbook choice of y and P *) - >> qabbrev_tac ‘pi = p3 ++ p2 ++ p1’ - >> qexistsl_tac [‘y’, ‘P’] >> art [] - >> NTAC 2 STRIP_TAC (* push ‘q <<= p’ to assumptions *) - (* RHS rewriting from M to M0 *) - >> Know ‘subterm X M0 q r = subterm X M q r’ - >- (qunabbrev_tac ‘M0’ \\ - MATCH_MP_TAC subterm_of_principle_hnf >> art []) - >> DISCH_THEN (ONCE_REWRITE_TAC o wrap o SYM) - (* LHS rewriting from M to M0 *) - >> Know ‘subterm X (apply pi M) q r = - subterm X (VAR b @* args' @* MAP VAR as) q r’ - >- (Q.PAT_X_ASSUM ‘_ = VAR b @* args' @* MAP VAR as’ - (ONCE_REWRITE_TAC o wrap o SYM) \\ - qabbrev_tac ‘t = apply pi M’ \\ - ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ - MATCH_MP_TAC subterm_of_principle_hnf >> art []) - >> Rewr' - (* stage cleanups *) - >> Q.PAT_X_ASSUM ‘solvable (apply pi M)’ K_TAC - >> Q.PAT_X_ASSUM ‘principle_hnf (apply pi M) = _’ K_TAC - >> Q.PAT_X_ASSUM ‘apply pi M == _’ K_TAC - >> Q.PAT_X_ASSUM ‘Boehm_transform pi’ K_TAC - (* stage work, now ‘M’ is eliminated from both sides! *) - >> Cases_on ‘q’ >- FULL_SIMP_TAC std_ss [] (* this asserts q = h::t *) - >> Know ‘h < m’ - >- (Cases_on ‘p’ >> fs [] \\ - Q.PAT_X_ASSUM ‘h = h'’ (fs o wrap o SYM) \\ - Know ‘subterm X M (h::t) r <> NONE’ - >- (FIRST_X_ASSUM MATCH_MP_TAC >> rw []) \\ - CONV_TAC (UNBETA_CONV “subterm X M (h::t) r”) \\ - qmatch_abbrev_tac ‘f _’ \\ - RW_TAC bool_ss [subterm_of_solvables] \\ - simp [Abbr ‘f’]) - >> DISCH_TAC - (* applying subterm_of_absfree_hnf *) - >> MP_TAC (Q.SPECL [‘X’, ‘VAR b @* args' @* MAP VAR as’, ‘h’, ‘t’, ‘r’] - subterm_of_absfree_hnf) - >> simp [hnf_appstar, GSYM appstar_APPEND, hnf_children_appstar] - >> DISCH_THEN K_TAC (* already used *) - (* eliminating ‘MAP VAR as’ *) - >> Know ‘EL h (args' ++ MAP VAR as) = EL h args'’ - >- (MATCH_MP_TAC EL_APPEND1 >> rw []) - >> Rewr' - (* eliminating ‘vs’ - - NOTE: ‘subterm Y’ changed to ‘subterm Z’ at next level. There's no - flexibility here on the choice of excluded variabes. - *) - >> Know ‘subterm X (LAMl vs (VAR y @* args)) (h::t) r = - subterm X (EL h args) t (SUC r)’ - >- (MP_TAC (Q.SPECL [‘X’, ‘LAMl vs (VAR y @* args)’, ‘h’, ‘t’, ‘r’] - subterm_of_hnf) \\ - simp [hnf_LAMl, hnf_appstar] \\ - DISCH_THEN K_TAC (* already used *) \\ - Q.PAT_X_ASSUM ‘M0 = LAMl vs (VAR y @* args)’ (REWRITE_TAC o wrap o SYM) \\ - simp [hnf_children_hnf]) - >> Rewr' - (* Now: subterm' Z (EL h args') t == [P/y] (subterm' Z (EL h args) t) - - First of all, those assumptions about p1,p2,p3 are no more needed. - *) - >> qunabbrev_tac ‘pi’ - >> Q.PAT_X_ASSUM ‘Boehm_transform p1’ K_TAC - >> Q.PAT_X_ASSUM ‘apply p1 M0 == M1’ K_TAC - >> qunabbrev_tac ‘p1’ - >> Q.PAT_X_ASSUM ‘Boehm_transform p2’ K_TAC - >> Q.PAT_X_ASSUM ‘apply p2 M1 = P @* args'’ K_TAC - >> qunabbrev_tac ‘p2’ - >> Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC - >> Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC - >> qunabbrev_tac ‘p3’ - >> Q.PAT_X_ASSUM ‘h::t <> []’ K_TAC - >> qabbrev_tac ‘N = EL h args’ - >> qabbrev_tac ‘N' = EL h args'’ - (* eliminating N' *) - >> ‘N' = [P/y] N’ by simp [EL_MAP, Abbr ‘m’, Abbr ‘N’, Abbr ‘N'’, Abbr ‘args'’] - >> POP_ORW - >> qunabbrev_tac ‘N'’ - (* cleanup args' *) - >> Q.PAT_X_ASSUM - ‘!i. i < m ==> - FV (EL i args') SUBSET FV (EL i args)’ K_TAC - >> Q.PAT_X_ASSUM ‘LENGTH args' = m’ K_TAC - >> qunabbrev_tac ‘args'’ - (* cleanup l, as and b *) - >> Q.PAT_X_ASSUM ‘ALL_DISTINCT l’ K_TAC - >> NTAC 2 (Q.PAT_X_ASSUM ‘DISJOINT (set l) _’ K_TAC) - >> Q.PAT_X_ASSUM ‘LENGTH l = q - m + 1’ K_TAC - >> Q.PAT_X_ASSUM ‘l <> []’ K_TAC - >> Q.PAT_X_ASSUM ‘l = SNOC b as’ K_TAC - >> Q.PAT_X_ASSUM ‘~MEM y l’ K_TAC - >> Q.PAT_X_ASSUM ‘LENGTH as = q - m’ K_TAC - >> qunabbrevl_tac [‘l’, ‘as’, ‘b’] - (* applying subterm_headvar_lemma *) - >> Know ‘hnf_headvar M1 IN X UNION RANK (SUC r)’ - >- (MATCH_MP_TAC subterm_headvar_lemma \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘vs’] >> simp []) - >> ASM_SIMP_TAC std_ss [hnf_head_hnf, THE_VAR_thm] - >> DISCH_TAC (* y IN X UNION RANK (SUC r) *) - (* applying subterm_subst_permutator_cong *) - >> MATCH_MP_TAC subterm_subst_permutator_cong' - >> Q.EXISTS_TAC ‘d’ - >> simp [Abbr ‘P’] - >> CONJ_TAC - >- (qunabbrev_tac ‘N’ >> MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) - (* stage work *) - >> CONJ_ASM1_TAC (* subterm Z N t <> NONE *) - >- (Q.PAT_X_ASSUM ‘!q. q <<= p ==> subterm X M q r <> NONE’ - (MP_TAC o Q.SPEC ‘h::t’) \\ - Q.PAT_X_ASSUM ‘M0 = _’ K_TAC \\ - simp [subterm_of_solvables]) - (* final goal: subterm_width (EL h args) t <= d *) - >> qunabbrev_tac ‘d’ - (* applying subterm_width_thm *) - >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_thm) - >> simp [] >> DISCH_THEN K_TAC - (* applying subterm_width_thm again *) - >> MP_TAC (Q.SPECL [‘X’, ‘N’, ‘t’, ‘SUC r’] subterm_width_thm) - >> Know ‘FV N SUBSET X UNION RANK (SUC r)’ - >- (qunabbrev_tac ‘N’ \\ + >- (Q.PAT_X_ASSUM ‘y IN X UNION RANK r’ MP_TAC \\ + Suff ‘X UNION RANK r SUBSET X UNION RANK (SUC r)’ + >- (REWRITE_TAC [SUBSET_DEF] \\ + DISCH_THEN (REWRITE_TAC o wrap)) \\ + Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [RANK_MONO]) \\ + qunabbrev_tac ‘N’ \\ MATCH_MP_TAC subterm_induction_lemma' \\ qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) >> DISCH_TAC - >> ‘t IN ltree_paths (BT' X N (SUC r))’ - by PROVE_TAC [subterm_imp_ltree_paths] - >> simp [] >> DISCH_THEN K_TAC - (* applying SUBSET_MAX_SET *) - >> MATCH_MP_TAC SUBSET_MAX_SET - >> CONJ_TAC - >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) + (* extra goal for subterm_width *) + >> qabbrev_tac ‘pi = ZIP (xs',ys')’ + >> Q.PAT_X_ASSUM ‘hnf t’ K_TAC + >> rfs [Abbr ‘t’, tpm_appstar] + >> qabbrev_tac ‘Ms = listpm term_pmact pi args''’ + >> Know ‘subterm_width ([P/y] M) (h::p) <= d <=> + d <= d /\ subterm_width (EL h Ms) p <= d’ + >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ + qexistsl_tac [‘X’, ‘r’, ‘M0'’, ‘n'’, ‘zs’, ‘M1'’] \\ + simp [principle_hnf_beta_reduce, ltree_paths_def] \\ + CONJ_TAC + >- (rw [FV_SUB, Abbr ‘P’, FV_permutator] \\ + MATCH_MP_TAC SUBSET_TRANS >> Q.EXISTS_TAC ‘FV M’ >> art [] \\ + SET_TAC []) \\ + simp [Abbr ‘M0'’]) + >> Rewr' + >> simp [Abbr ‘Ms’, EL_MAP] + >> ‘EL h args'' = EL h args'’ by rw [Abbr ‘args''’, EL_APPEND1] + >> POP_ORW + >> qunabbrev_tac ‘pi’ + >> Know ‘tpm (ZIP (xs',ys')) (EL h args') = EL h args'’ + >- (MATCH_MP_TAC tpm_unchanged >> simp [MAP_ZIP] \\ + ONCE_REWRITE_TAC [DISJOINT_SYM] \\ + CONJ_TAC \\ (* 2 subgoals, same tactics *) + FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) + >> Rewr' + >> simp [Abbr ‘args'’, EL_MAP] + >> FIRST_X_ASSUM (MATCH_MP_TAC o cj 2) >> art [] + >> Q.EXISTS_TAC ‘SUC r’ >> art [] >> CONJ_TAC - >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) - (* final goal *) - >> ‘hnf_children_size M0 = m’ by rw [Abbr ‘m’] - >> Q.PAT_X_ASSUM ‘M0 = _’ K_TAC - >> rw [SUBSET_DEF] (* this asserts q <<= t *) - >> Know ‘h::q <<= p’ - >- (MATCH_MP_TAC IS_PREFIX_TRANS \\ - Q.EXISTS_TAC ‘h::t’ >> simp []) - >> DISCH_TAC - >> Q.EXISTS_TAC ‘h::q’ >> simp [] - >> Know ‘subterm X M (h::q) r <> NONE’ - >- (FIRST_X_ASSUM MATCH_MP_TAC >> simp []) - >> Know ‘subterm X N q (SUC r) <> NONE’ - >- (Cases_on ‘t = []’ >- fs [] \\ - irule (cj 1 subterm_solvable_lemma) >> simp [] \\ - Q.EXISTS_TAC ‘t’ >> art []) - >> DISCH_TAC - >> Q_TAC (UNBETA_TAC [subterm_of_solvables]) ‘subterm' X M (h::q) r’ - >> simp [] + >- (qunabbrev_tac ‘N’ \\ + MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) + >> Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC + >> qunabbrev_tac ‘Y’ + >> Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ + >- METIS_TAC [SUBSET_DEF] + >> Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] + >> rw [RANK_MONO] QED -(* Proposition 10.3.7 (i) [1, p.248] (Boehm out lemma) - - NOTE: This time the case ‘p = []’ can be included, but it's a trvial case. - - NOTE: The original lemma in textbook requires ‘p IN ltree_paths (BT X M)’, - but this seems wrong, as ‘subterm X M p’ may not be defined if only ‘p’ is - valid path (i.e. the subterm could be a bottom (\bot) as the result of un- - solvable terms). - - NOTE: Can we enhance this lemma by using ‘-h->*’ instead of ‘==’? - - NOTE: Use subterm_imp_ltree_paths to prove ‘p IN ltree_paths (BT X M)’. - *) -Theorem Boehm_out_lemma : - !X p M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - subterm X M p r <> NONE ==> - ?pi. Boehm_transform pi /\ - ?ss. apply pi M == subterm' X M p r ISUB ss +Theorem subterm_width_isub_permutator_cong : + !ys p X M r P d ss. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + P = permutator d /\ + (!y. MEM y ys ==> y IN X UNION RANK r) /\ + ss = MAP (\y. (P,y)) ys /\ + p IN ltree_paths (BT' X M r) /\ + subterm_width M p <= d + ==> p IN ltree_paths (BT' X (M ISUB ss) r) /\ + subterm_width (M ISUB ss) p <= d Proof - Q.X_GEN_TAC ‘X’ - >> Induct_on ‘p’ - >- (rw [] \\ - Q.EXISTS_TAC ‘[]’ >> rw [] \\ - Q.EXISTS_TAC ‘[]’ >> rw []) - >> rpt STRIP_TAC - >> rename1 ‘subterm X M (h::t) r <> NONE’ - >> qabbrev_tac ‘p = h::t’ (* head and tail *) - >> ‘p <> []’ by rw [Abbr ‘p’] - >> ‘(!q. q <<= p ==> subterm X M q r <> NONE) /\ - !q. q <<= FRONT p ==> solvable (subterm' X M q r)’ - by METIS_TAC [subterm_solvable_lemma] - >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘h::t’, ‘r’] Boehm_transform_exists_lemma) + Induct_on ‘ys’ >- rw [] + >> rpt GEN_TAC + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘ss = _’ (REWRITE_TAC o wrap) + >> simp [] + >> Q.PAT_X_ASSUM ‘P = permutator d’ K_TAC + >> qabbrev_tac ‘P = permutator d’ + >> qabbrev_tac ‘N = [P/h] M’ + >> qabbrev_tac ‘ss = MAP (\y. (P,y)) ys’ + >> MP_TAC (Q.SPECL [‘X’, ‘P’, ‘d’, ‘h’, ‘p’, ‘M’, ‘r’] + subterm_width_subst_permutator_cong) >> simp [] - >> DISCH_THEN (Q.X_CHOOSE_THEN ‘p0’ MP_TAC) - >> RW_TAC std_ss [] (* push p0 properties to assumptions *) - >> POP_ASSUM (MP_TAC o (Q.SPEC ‘p’)) (* put q = p *) - >> rw [] - >> qabbrev_tac ‘M' = apply p0 M’ - >> ‘solvable M' /\ ?y Ms. M' -h->* VAR y @* Ms /\ EVERY (\e. y # e) Ms’ - by METIS_TAC [is_ready_alt'] - >> ‘principle_hnf M' = VAR y @* Ms’ by rw [principle_hnf_thm', hnf_appstar] - (* stage work *) - >> qunabbrev_tac ‘p’ - >> Know ‘h < LENGTH Ms’ - >- (Q.PAT_X_ASSUM ‘subterm X M' (h::t) r <> NONE’ MP_TAC \\ - RW_TAC std_ss [subterm_of_solvables] >> fs []) - >> DISCH_TAC - >> qabbrev_tac ‘m = LENGTH Ms’ - >> qabbrev_tac ‘N = EL h Ms’ - (* stage work *) - >> Know ‘subterm X N t (SUC r) <> NONE /\ - subterm' X M' (h::t) r = subterm' X N t (SUC r)’ - >- (Q.PAT_X_ASSUM ‘subterm' X M' (h::t) r = - [P/v] (subterm' X M (h::t) r)’ K_TAC \\ - Q.PAT_X_ASSUM ‘subterm X M' (h::t) r <> NONE’ MP_TAC \\ - rw [subterm_of_solvables, Abbr ‘N’]) >> STRIP_TAC - >> Q.PAT_X_ASSUM ‘subterm' X M' (h::t) r = subterm' X N t (SUC r)’ (fs o wrap) - >> T_TAC - (* stage work, now define a selector *) - >> qabbrev_tac ‘U = selector h m’ - >> qabbrev_tac ‘p1 = [[U/y]]’ - >> ‘Boehm_transform p1’ by rw [Abbr ‘p1’] - >> qabbrev_tac ‘p10 = p1 ++ p0’ - >> ‘Boehm_transform p10’ by rw [Abbr ‘p10’, Boehm_transform_APPEND] - (* applying properties of selector (U) *) - >> Know ‘apply p10 M == N’ - >- (rw [Abbr ‘p10’, Boehm_apply_APPEND] \\ - MATCH_MP_TAC lameq_TRANS \\ - Q.EXISTS_TAC ‘apply p1 (principle_hnf M')’ \\ - CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong \\ - CONJ_TAC >- art [] \\ - MATCH_MP_TAC lameq_SYM \\ - MATCH_MP_TAC lameq_principle_hnf' >> art []) \\ - rw [Abbr ‘p1’, appstar_SUB] \\ - Know ‘MAP [U/y] Ms = Ms’ - >- (rw [LIST_EQ_REWRITE, EL_MAP] \\ - MATCH_MP_TAC lemma14b \\ - Q.PAT_X_ASSUM ‘EVERY (\e. y # e) Ms’ MP_TAC \\ - rw [EVERY_MEM, MEM_EL] \\ - POP_ASSUM MATCH_MP_TAC >> rename1 ‘i < m’ \\ - Q.EXISTS_TAC ‘i’ >> art []) >> Rewr' \\ - qunabbrevl_tac [‘U’, ‘N’] \\ - MATCH_MP_TAC selector_thm >> rw [Abbr ‘m’]) - >> DISCH_TAC - (* stage work, now using IH *) - >> Q.PAT_X_ASSUM ‘!M r. _’ (MP_TAC o (Q.SPECL [‘N’, ‘SUC r’])) >> simp [] - >> Know ‘FV N SUBSET X UNION RANK (SUC r)’ - >- (qunabbrevl_tac [‘N’, ‘M'’] \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘FV (apply p0 M)’ >> art [] \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘FV (VAR y @* Ms)’ \\ - reverse CONJ_TAC >- (MATCH_MP_TAC hreduce_FV_SUBSET >> art []) \\ - rw [SUBSET_DEF, FV_appstar, IN_UNION] \\ - DISJ2_TAC \\ - Q.EXISTS_TAC ‘FV (EL h Ms)’ >> art [] \\ - Q.EXISTS_TAC ‘EL h Ms’ >> rw [EL_MEM]) - >> RW_TAC std_ss [] - >> rename1 ‘apply p2 N == _ ISUB ss'’ - >> qabbrev_tac ‘N' = subterm' X M (h::t) r’ - (* final stage *) - >> Q.EXISTS_TAC ‘p2 ++ p10’ - >> CONJ_TAC - >- (MATCH_MP_TAC Boehm_transform_APPEND >> art []) - >> Q.EXISTS_TAC ‘[(P,v)] ++ ss'’ - >> MATCH_MP_TAC lameq_TRANS - >> Q.EXISTS_TAC ‘apply p2 N’ - >> simp [ISUB_def] - >> rw [Boehm_apply_APPEND] - >> MATCH_MP_TAC Boehm_apply_lameq_cong >> art [] + >> FIRST_X_ASSUM MATCH_MP_TAC + >> Q.EXISTS_TAC ‘P’ >> simp [] + >> Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] + >> rw [FV_SUB, Abbr ‘N’] + >> rw [Abbr ‘P’, FV_permutator] QED -(*---------------------------------------------------------------------------* - * Faithfulness and agreements of terms - *---------------------------------------------------------------------------*) - -(* Definition 10.2.21 (i) [1, p.238] - - NOTE: For ‘y1 = y2’ to be meaningful, here we assumed that vs1 and vs2 - share the same prefix, i.e. either vs1 <<= vs2 or vs2 <<= vs1. In reality, - we have ‘vs1 = RNEWS r n1 X /\ vs2 = RNEWS r n2 X’ for some X and r. - *) -Definition head_equivalent_def : - head_equivalent ((a1,m1) :BT_node # num option) - ((a2,m2) :BT_node # num option) = - case (a1,a2) of - (SOME (vs1,y1),SOME (vs2,y2)) => - y1 = y2 /\ LENGTH vs1 + THE m2 = LENGTH vs2 + THE m1 - | (SOME _,NONE) => F - | (NONE,SOME _) => F - | (NONE,NONE) => T -End - -(* Alternative version without ‘y1 = y2’ *) -Definition head_equivalent'_def : - head_equivalent' ((a1,m1) :BT_node # num option) - ((a2,m2) :BT_node # num option) = - case (a1,a2) of - (SOME (vs1,_),SOME (vs2,_)) => - LENGTH vs1 + THE m2 = LENGTH vs2 + THE m1 - | (SOME _,NONE) => F - | (NONE,SOME _) => F - | (NONE,NONE) => T -End - -Theorem head_equivalent_imp_head_equivalent' : - !A B. head_equivalent A B ==> head_equivalent' A B +Theorem subterm_width_isub_permutator_cong_alt : + !X p r d y k ss M. + FINITE X /\ FV M SUBSET X UNION RANK r /\ + (!i. i < k ==> y i IN X UNION RANK r) /\ + ss = GENLIST (\i. (permutator (d + i),y i)) k /\ + p IN ltree_paths (BT' X M r) /\ + subterm_width M p <= d + ==> p IN ltree_paths (BT' X (M ISUB ss) r) /\ + subterm_width (M ISUB ss) p <= d + k Proof - rpt GEN_TAC - >> Cases_on ‘A’ >> Cases_on ‘B’ - >> rw [head_equivalent_def, head_equivalent'_def] - >> Cases_on ‘q’ >> fs [] - >> Cases_on ‘q'’ >> fs [] - >> Cases_on ‘x'’ >> fs [] - >> Cases_on ‘x’ >> fs [] + qx_genl_tac [‘X’, ‘p’, ‘r’, ‘d’, ‘y’] + >> Induct_on ‘k’ >- rw [] + >> qx_genl_tac [‘ss'’, ‘M’] + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘ss' = _’ (REWRITE_TAC o wrap) + >> SIMP_TAC std_ss [GENLIST, ISUB_SNOC] + >> qabbrev_tac ‘P = \i. permutator (d + i)’ >> fs [] + >> qabbrev_tac ‘ss = GENLIST (\i. (P i,y i)) k’ + >> Q.PAT_X_ASSUM ‘!M'. FV M' SUBSET X UNION RANK r /\ _ ==> _’ + (MP_TAC o Q.SPEC ‘M’) + >> simp [] + >> STRIP_TAC + >> qabbrev_tac ‘N = M ISUB ss’ + >> qabbrev_tac ‘Q = P k’ + >> qabbrev_tac ‘v = y k’ + >> qabbrev_tac ‘w = d + k’ + >> MP_TAC (Q.SPECL [‘X’, ‘Q’, ‘w’, ‘v’, ‘p’, ‘N’, ‘r’] + subterm_width_subst_permutator_cong) + >> simp [Abbr ‘Q’, Abbr ‘v’, Abbr ‘w’] + >> impl_tac + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] \\ + qunabbrev_tac ‘N’ \\ + MP_TAC (Q.SPECL [‘ss’, ‘M’] FV_ISUB_upperbound) \\ + Suff ‘FVS ss = {}’ >- simp [] \\ + simp [Abbr ‘ss’, FVS_ALT] \\ + Cases_on ‘k = 0’ >> simp [] \\ + DISJ2_TAC \\ + simp [MAP_GENLIST, LIST_TO_SET_GENLIST] \\ + simp [Abbr ‘P’, FV_permutator, o_DEF] \\ + simp [IMAGE_CONST]) + >> rw [] QED -Theorem head_equivalent_refl[simp] : - head_equivalent A A -Proof - Cases_on ‘A’ >> rw [head_equivalent_def] - >> Cases_on ‘q’ >> rw [] - >> Cases_on ‘x’ >> rw [] -QED +(* This theorem is based on FV_subterm_lemma -Theorem head_equivalent'_refl[simp] : - head_equivalent' A A + NOTE: If initially ‘FV M SUBSET X UNION RANK r’ and ‘v # M’ but + ‘v IN RANK r’, each deeper subterm only adds FV from ROW r onwards, + thus ‘v # subterm X M p r’ still holds. + *) +Theorem FV_subterm_thm : + !X v p M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm X M p r <> NONE /\ + v # M /\ v IN X UNION RANK r ==> v # (subterm' X M p r) Proof - Cases_on ‘A’ >> rw [head_equivalent'_def] - >> Cases_on ‘q’ >> rw [] - >> Cases_on ‘x’ >> rw [] + NTAC 2 GEN_TAC + >> Induct_on ‘p’ >- rw [] + >> rpt GEN_TAC + >> STRIP_TAC + >> Know ‘(!q. q <<= h::p ==> subterm X M q r <> NONE) /\ + (!q. q <<= FRONT (h::p) ==> solvable (subterm' X M q r))’ + >- (MATCH_MP_TAC subterm_solvable_lemma >> simp []) + >> STRIP_TAC + >> Know ‘solvable M’ + >- (Q.PAT_X_ASSUM ‘!q. q <<= FRONT (h::p) ==> solvable _’ + (MP_TAC o (Q.SPEC ‘[]’)) >> rw []) + >> DISCH_TAC + >> Q.PAT_X_ASSUM ‘subterm X M (h::p) r <> NONE’ MP_TAC + >> Q_TAC (UNBETA_TAC [subterm_alt]) ‘subterm X M (h::p) r’ + >> STRIP_TAC + >> FIRST_X_ASSUM MATCH_MP_TAC >> art [] + >> CONJ_TAC + >- (MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) + >> reverse CONJ_TAC + >- (Q.PAT_X_ASSUM ‘v IN X UNION RANK r’ MP_TAC \\ + Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [RANK_MONO]) + >> Know ‘LENGTH Ms = m’ + >- (simp [Once EQ_SYM_EQ, Abbr ‘m’] \\ + MATCH_MP_TAC hnf_children_size_alt \\ + qexistsl_tac [‘X’, ‘M’, ‘r’, ‘n’, ‘vs’, ‘M1’] >> simp []) + >> DISCH_TAC + (* applying FV_subterm_lemma *) + >> Know ‘FV (EL h Ms) SUBSET FV M UNION set vs’ + >- (MATCH_MP_TAC FV_subterm_lemma \\ + qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘m’, ‘M1’] >> simp []) + >> DISCH_TAC + >> STRIP_TAC + >> Know ‘v IN FV M UNION set vs’ >- METIS_TAC [SUBSET_DEF] + >> rw [IN_UNION] + >> Q.PAT_X_ASSUM ‘v IN X UNION RANK r’ MP_TAC + >> rw [IN_UNION] + >- (qunabbrev_tac ‘vs’ \\ + Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ \\ + Q.PAT_X_ASSUM ‘DISJOINT (set vs) X’ MP_TAC \\ + rw [DISJOINT_ALT']) + >> Suff ‘DISJOINT (RANK r) (set vs)’ >- rw [DISJOINT_ALT] + >> qunabbrev_tac ‘vs’ + >> MATCH_MP_TAC DISJOINT_RANK_RNEWS >> simp [] QED -Theorem head_equivalent_sym : - !A B. head_equivalent A B ==> head_equivalent B A +Theorem subterm_once_fresh_tpm_cong[local] : + !p X M r v v'. FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm X M p r <> NONE /\ + v IN X UNION RANK r /\ + v' IN X UNION RANK r /\ v' # M + ==> subterm X (tpm [(v,v')] M) p r <> NONE /\ + subterm' X (tpm [(v,v')] M) p r = tpm [(v,v')] (subterm' X M p r) Proof - qx_genl_tac [‘A’, ‘B’] - >> Cases_on ‘A’ >> Cases_on ‘B’ >> simp [head_equivalent_def] - >> Cases_on ‘q’ >> Cases_on ‘q'’ >> simp [] - >> Cases_on ‘x’ >> Cases_on ‘x'’ >> simp [] + rpt GEN_TAC >> STRIP_TAC + >> Know ‘v' # subterm' X M p r’ + >- (MATCH_MP_TAC FV_subterm_thm >> art []) + >> DISCH_TAC + >> simp [fresh_tpm_subst'] + >> MATCH_MP_TAC subterm_fresh_subst_cong >> art [] QED -Theorem head_equivalent_comm : - !A B. head_equivalent A B <=> head_equivalent B A +Theorem subterm_fresh_tpm_cong_lemma[local] : + !pi X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm X M p r <> NONE /\ + set (MAP FST pi) SUBSET X UNION RANK r /\ + set (MAP SND pi) SUBSET X UNION RANK r /\ + ALL_DISTINCT (MAP FST pi) /\ + ALL_DISTINCT (MAP SND pi) /\ + DISJOINT (set (MAP FST pi)) (set (MAP SND pi)) /\ + DISJOINT (set (MAP SND pi)) (FV M) + ==> subterm X (tpm pi M) p r <> NONE /\ + subterm' X (tpm pi M) p r = tpm pi (subterm' X M p r) Proof - rpt GEN_TAC - >> EQ_TAC >> rw [head_equivalent_sym] + Induct_on ‘pi’ >- rw [] + >> simp [FORALL_PROD] + >> qx_genl_tac [‘u’, ‘v’] + >> rpt GEN_TAC + >> STRIP_TAC + >> ‘!t. tpm ((u,v)::pi) t = tpm [(u,v)] (tpm pi t)’ by rw [Once tpm_CONS] + >> POP_ORW + >> qabbrev_tac ‘N = tpm pi M’ + (* applying IH *) + >> Q.PAT_X_ASSUM ‘!X M p r. P’ (MP_TAC o Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’]) + >> simp [] + >> STRIP_TAC + >> POP_ASSUM (REWRITE_TAC o wrap o SYM) + >> MATCH_MP_TAC subterm_once_fresh_tpm_cong + >> simp [] + >> qabbrev_tac ‘vs = MAP FST pi’ + >> qabbrev_tac ‘vs' = MAP SND pi’ + >> ‘LENGTH vs' = LENGTH vs’ by rw [Abbr ‘vs’, Abbr ‘vs'’] + >> CONJ_TAC + >- (simp [Abbr ‘N’, SUBSET_DEF, FV_tpm] \\ + rpt STRIP_TAC \\ + qabbrev_tac ‘y = lswapstr (REVERSE pi) x’ \\ + ‘x = lswapstr pi y’ by rw [Abbr ‘y’] >> POP_ORW \\ + Know ‘pi = ZIP (vs,vs')’ + >- (simp [Abbr ‘vs’, Abbr ‘vs'’, LIST_EQ_REWRITE] \\ + rw [EL_ZIP, EL_MAP]) >> Rewr' \\ + Cases_on ‘MEM y vs’ + >- (Suff ‘lswapstr (ZIP (vs,vs')) y IN set vs'’ >- ASM_SET_TAC [] \\ + MATCH_MP_TAC MEM_lswapstr >> simp []) \\ + Cases_on ‘MEM y vs'’ + >- (Suff ‘lswapstr (ZIP (vs,vs')) y IN set vs’ >- ASM_SET_TAC [] \\ + MATCH_MP_TAC MEM_lswapstr' >> simp []) \\ + Suff ‘lswapstr (ZIP (vs,vs')) y = y’ >- ASM_SET_TAC [] \\ + MATCH_MP_TAC lswapstr_unchanged' \\ + simp [MAP_ZIP]) + >> simp [Abbr ‘N’] + >> Suff ‘lswapstr (REVERSE pi) v = v’ >- rw [] + >> MATCH_MP_TAC lswapstr_unchanged' + >> simp [MAP_REVERSE, MEM_REVERSE] QED -(* Definition 10.2.21 (ii) [1, p.238] *) -Overload ltree_equiv = “OPTREL head_equivalent” -Overload ltree_equiv' = “OPTREL head_equivalent'” - -Theorem ltree_equiv_imp_ltree_equiv' : - !A B. ltree_equiv A B ==> ltree_equiv' A B +Theorem subterm_fresh_tpm_cong : + !pi X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + set (MAP FST pi) SUBSET X UNION RANK r /\ + set (MAP SND pi) SUBSET X UNION RANK r /\ + ALL_DISTINCT (MAP FST pi) /\ + ALL_DISTINCT (MAP SND pi) /\ + DISJOINT (set (MAP FST pi)) (set (MAP SND pi)) /\ + DISJOINT (set (MAP SND pi)) (FV M) + ==> (subterm X M p r = NONE <=> subterm X (tpm pi M) p r = NONE) /\ + (subterm X M p r <> NONE ==> + subterm' X (tpm pi M) p r = tpm pi (subterm' X M p r)) Proof - rpt GEN_TAC - >> Cases_on ‘A’ >> Cases_on ‘B’ >> rw [] - >> MATCH_MP_TAC head_equivalent_imp_head_equivalent' >> art [] + reverse (rpt STRIP_TAC) + >- (MATCH_MP_TAC (cj 2 subterm_fresh_tpm_cong_lemma) >> art []) + >> reverse EQ_TAC + >- (CCONTR_TAC >> fs [] \\ + Q.PAT_X_ASSUM ‘subterm X (tpm pi M) p r = NONE’ MP_TAC \\ + simp [] \\ + MATCH_MP_TAC (cj 1 subterm_fresh_tpm_cong_lemma) >> art []) +(* stage work *) + >> CCONTR_TAC >> fs [] + >> qabbrev_tac ‘N = tpm pi M’ + >> Q.PAT_X_ASSUM ‘subterm X M p r = NONE’ MP_TAC + >> ‘M = tpm (REVERSE pi) N’ by rw [Abbr ‘N’] + >> simp [] + >> qabbrev_tac ‘xs = MAP FST pi’ + >> qabbrev_tac ‘ys = MAP SND pi’ + >> ‘LENGTH ys = LENGTH xs’ by rw [Abbr ‘xs’, Abbr ‘ys’] + >> Know ‘pi = ZIP (xs,ys)’ + >- (rw [Abbr ‘xs’, Abbr ‘ys’, ZIP_MAP] \\ + rw [LIST_EQ_REWRITE, EL_MAP]) + >> DISCH_TAC + >> simp [REVERSE_ZIP] + >> qabbrev_tac ‘xs' = REVERSE xs’ + >> qabbrev_tac ‘ys' = REVERSE ys’ + >> ‘LENGTH ys' = LENGTH xs'’ by rw [Abbr ‘xs'’, Abbr ‘ys'’] + (* applying pmact_flip_args_all *) + >> ‘tpm (ZIP (xs',ys')) N = tpm (ZIP (ys',xs')) N’ by rw [pmact_flip_args_all] + >> POP_ORW + >> qabbrev_tac ‘pi' = ZIP (ys',xs')’ + >> MATCH_MP_TAC (cj 1 subterm_fresh_tpm_cong_lemma) + >> simp [Abbr ‘pi'’, MAP_ZIP] + (* applying FV_tpm_lemma *) + >> CONJ_TAC + >- (qunabbrev_tac ‘N’ \\ + MATCH_MP_TAC FV_tpm_lemma \\ + Q.EXISTS_TAC ‘r’ >> simp [MAP_ZIP] \\ + ASM_SET_TAC []) + >> CONJ_TAC + >- (Q.PAT_X_ASSUM ‘set ys SUBSET X UNION RANK r’ MP_TAC \\ + rw [SUBSET_DEF, Abbr ‘ys'’]) + >> CONJ_TAC + >- (Q.PAT_X_ASSUM ‘set xs SUBSET X UNION RANK r’ MP_TAC \\ + rw [SUBSET_DEF, Abbr ‘xs'’]) + >> CONJ_TAC >- rw [ALL_DISTINCT_REVERSE, Abbr ‘ys'’] + >> CONJ_TAC >- rw [ALL_DISTINCT_REVERSE, Abbr ‘xs'’] + >> CONJ_TAC + >- (rw [DISJOINT_ALT', Abbr ‘xs'’] \\ + Suff ‘DISJOINT (set xs) (set ys')’ >- rw [DISJOINT_ALT] \\ + rw [DISJOINT_ALT', Abbr ‘ys'’] \\ + Q.PAT_X_ASSUM ‘DISJOINT (set xs) (set ys)’ MP_TAC \\ + rw [DISJOINT_ALT']) + (* stage work *) + >> Suff ‘DISJOINT (set xs) (FV N)’ + >- rw [DISJOINT_ALT, Abbr ‘xs'’] + >> POP_ASSUM K_TAC + >> qunabbrevl_tac [‘xs'’, ‘ys'’] + >> simp [Abbr ‘N’] + >> POP_ASSUM K_TAC (* pi = ZIP (xs,ys) *) + >> ONCE_REWRITE_TAC [DISJOINT_SYM] + >> Know ‘tpm (ZIP (xs,ys)) M = M ISUB ZIP (MAP VAR ys,xs)’ + >- (MATCH_MP_TAC fresh_tpm_isub' >> art []) + >> Rewr' + >> MATCH_MP_TAC FV_renaming_disjoint >> art [] QED -Theorem ltree_equiv_refl[simp] : - ltree_equiv A A -Proof - MATCH_MP_TAC OPTREL_refl >> rw [] -QED +(* Lemma 10.3.6 (ii) [1, p.247]: -Theorem ltree_equiv_sym : - !A B. ltree_equiv A B ==> ltree_equiv B A -Proof - rpt GEN_TAC - >> Cases_on ‘A’ >> Cases_on ‘B’ >> rw [OPTREL_THM] - >> rw [Once head_equivalent_comm] -QED + NOTE: The construction of ‘pi’ needs a fixed ltree path ‘p’, so that we can + collect the maximum number of children in all nodes along ‘p’. In another + word, there exists no universal ‘pi’ for which the conclusion holds for + arbitrary ‘p’. -Theorem ltree_equiv_comm : - !A B. ltree_equiv A B <=> ltree_equiv B A -Proof - rpt STRIP_TAC - >> EQ_TAC >> rw [ltree_equiv_sym] -QED + NOTE: Added ‘subterm X M p <> NONE’ to antecedents so that ‘subterm' X M p’ + is defined/specified. ‘subterm X (apply pi M) p <> NONE’ can be derived. -Theorem ltree_equiv_some_bot_imp : - !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - ltree_equiv (SOME bot) (ltree_el (BT' X M r) p) ==> - ltree_el (BT' X M r) p = SOME bot -Proof - rw [OPTREL_def] - >> Cases_on ‘y0’ >> fs [head_equivalent_def] - >> Cases_on ‘q’ >> fs [] - >> METIS_TAC [BT_ltree_el_eq_some_none] -QED + NOTE: ‘p <> []’ must be added into antecedents, otherwise the statement + becomes: -(* |- !X M p r. - FINITE X /\ FV M SUBSET X UNION RANK r /\ - ltree_equiv (ltree_el (BT' X M r) p) (SOME bot) ==> - ltree_el (BT' X M r) p = SOME bot - *) -Theorem ltree_equiv_some_bot_imp' = - ONCE_REWRITE_RULE [ltree_equiv_comm] ltree_equiv_some_bot_imp + [...] |- !X M. ?pi. Boehm_transform pi /\ is_ready (apply pi M) /\ + ?fm. apply pi M == fm ' M -(* Definition 10.2.32 (v) [1, p.245] *) -Definition subtree_equiv_def : - subtree_equiv X M N p r = - ltree_equiv (ltree_el (BT' X M r) p) (ltree_el (BT' X N r) p) -End + which is impossible if M is not already "is_ready". -Definition subtree_equiv'_def : - subtree_equiv' X M N p r = - ltree_equiv' (ltree_el (BT' X M r) p) (ltree_el (BT' X N r) p) -End + NOTE: Added ‘p IN ltree_paths (BT X (apply pi M))’ into conclusions for the + needs in the user theorem. -(* NOTE: This theorem is possible but not easy -Theorem equivalent_alt_ltree_equiv_NIL : - !M N. FV M UNION FV N SUBSET X UNION RANK r ==> - (equivalent M N <=> subtree_equiv X M N [] r) -Proof - ... -QED - *) + NOTE: Extended the conclusion with ‘!q. q <<= p /\ q <> []’ -(* NOTE: ‘ltree_paths (BT' X M r) SUBSET ltree_paths (BT' X ([P/v] M) r)’ doesn't - hold. Instead, we need to consider certain p and ‘d <= subterm_width M p’. - This theorem holds even when M is not solvable. + NOTE: ‘FV (apply pi M) SUBSET X UNION RANK (SUC r)’ is added into the + conclusions for the needs of Boehm_out_lemma. *) -Theorem BT_ltree_paths_subst_cong : - !X P d v p M r. - FINITE X /\ FV M SUBSET X UNION RANK r /\ v IN X UNION RANK r /\ - P = permutator d /\ subterm_width M p <= d /\ - p IN ltree_paths (BT' X M r) ==> - p IN ltree_paths (BT' X ([P/v] M) r) /\ - subterm_width ([P/v] M) p <= d +Theorem Boehm_transform_exists_lemma : + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + p <> [] /\ subterm X M p r <> NONE ==> + ?pi. Boehm_transform pi /\ + solvable (apply pi M) /\ is_ready (apply pi M) /\ + FV (apply pi M) SUBSET X UNION RANK (SUC r) /\ + ?v P. closed P /\ + !q. q <<= p /\ q <> [] ==> + subterm X (apply pi M) q r <> NONE /\ + subterm' X (apply pi M) q r = [P/v] (subterm' X M q r) Proof - simp [ltree_paths_def] - >> NTAC 4 GEN_TAC - >> Induct_on ‘p’ - >- (rw [ltree_lookup] \\ - qabbrev_tac ‘P = permutator d’ \\ - MATCH_MP_TAC (cj 2 subterm_subst_permutator_cong) \\ - qexistsl_tac [‘X’, ‘r’] >> simp []) - >> rpt GEN_TAC >> STRIP_TAC - >> qabbrev_tac ‘P = permutator d’ - >> ‘h::p IN ltree_paths (BT' X M r)’ by rw [ltree_paths_def] - >> Q.PAT_X_ASSUM ‘ltree_lookup (BT' X M r) (h::p) <> NONE’ MP_TAC - >> qabbrev_tac ‘Y = X UNION RANK r’ - >> reverse (Cases_on ‘solvable M’) - >- simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def] + rpt STRIP_TAC + >> ‘p IN ltree_paths (BT' X M r)’ by PROVE_TAC [subterm_imp_ltree_paths] + >> ‘(!q. q <<= p ==> subterm X M q r <> NONE) /\ + !q. q <<= FRONT p ==> solvable (subterm' X M q r)’ + by (MATCH_MP_TAC subterm_solvable_lemma >> art []) + >> Know ‘solvable M’ + >- (POP_ASSUM (MP_TAC o Q.SPEC ‘[]’) >> rw []) >> DISCH_TAC + (* M0 is now meaningful since M is solvable *) >> qabbrev_tac ‘M0 = principle_hnf M’ + >> ‘hnf M0’ by PROVE_TAC [hnf_principle_hnf'] >> qabbrev_tac ‘n = LAMl_size M0’ + (* NOTE: here the excluded list must contain ‘FV M’. Just ‘FV M0’ doesn't + work later, when calling the important [principle_hnf_denude_thm]. + This is conflicting with BT_generator_def and subterm_def. + *) >> Q_TAC (RNEWS_TAC (“vs :string list”, “r :num”, “n :num”)) ‘X’ >> ‘DISJOINT (set vs) (FV M)’ by METIS_TAC [subterm_disjoint_lemma] - >> ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] - >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ - >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, - “y :string”, “args :term list”)) ‘M1’ + >> Know ‘?y args. M0 = LAMl (TAKE n vs) (VAR y @* args)’ + >- (qunabbrev_tac ‘n’ \\ + ‘DISJOINT (set vs) (FV M0)’ by METIS_TAC [subterm_disjoint_lemma'] \\ + irule (iffLR hnf_cases_shared) >> rw [] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘FV M’ >> art [] \\ + qunabbrev_tac ‘M0’ >> MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) + >> STRIP_TAC + (* eliminate ‘TAKE n vs’ *) >> ‘TAKE n vs = vs’ by rw [] - >> POP_ASSUM (rfs o wrap) - >> Q.PAT_X_ASSUM ‘DISJOINT (set vs) (FV M0)’ K_TAC - >> Know ‘~MEM v vs’ - >- (Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ - rw [Abbr ‘Y’, IN_UNION] - >- (Q.PAT_X_ASSUM ‘DISJOINT (set vs) X’ MP_TAC \\ - rw [DISJOINT_ALT']) \\ - Suff ‘DISJOINT (RANK r) (set vs)’ >- rw [DISJOINT_ALT] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC DISJOINT_RANK_RNEWS' >> art []) - >> DISCH_TAC - >> Know ‘solvable ([P/v] M)’ - >- (MATCH_MP_TAC solvable_subst_permutator \\ - qexistsl_tac [‘X’, ‘r’, ‘d’] >> simp [] \\ - MP_TAC (Q.SPECL [‘X’, ‘M’, ‘h::p’, ‘r’] subterm_width_first) \\ - simp [ltree_paths_def]) + >> POP_ASSUM (REV_FULL_SIMP_TAC std_ss o wrap) + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vs)’ + >> Know ‘M1 = VAR y @* args’ + >- (qunabbrev_tac ‘M1’ >> POP_ORW \\ + MATCH_MP_TAC principle_hnf_beta_reduce >> rw [hnf_appstar]) >> DISCH_TAC - >> ‘M -h->* M0’ by METIS_TAC [principle_hnf_thm'] - >> Know ‘[P/v] M -h->* [P/v] M0’ >- PROVE_TAC [hreduce_substitutive] - >> ‘DISJOINT (set vs) (FV P)’ by rw [DISJOINT_ALT', FV_permutator, Abbr ‘P’] - >> simp [LAMl_SUB, appstar_SUB] - >> POP_ASSUM K_TAC (* DISJOINT (set vs) (FV P) *) - >> qabbrev_tac ‘args' = MAP [P/v] args’ >> qabbrev_tac ‘m = LENGTH args’ - >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’, Abbr ‘m’] + (* using ‘subterm_width’ and applying subterm_width_thm *) + >> qabbrev_tac ‘d = subterm_width M p’ + >> Know ‘m <= d’ + >- (MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_first) \\ + rw [Abbr ‘d’]) >> DISCH_TAC + (* p1 is the first Boehm transformation for removing abstractions of M0 *) + >> qabbrev_tac ‘p1 = MAP rightctxt (REVERSE (MAP VAR vs))’ + >> ‘Boehm_transform p1’ by rw [Abbr ‘p1’, MAP_MAP_o, GSYM MAP_REVERSE] + >> ‘apply p1 M0 == M1’ by rw [Abbr ‘p1’, Boehm_apply_MAP_rightctxt'] + (* stage work (all correct until here) + + Now we define the permutator P (and then p2). This requires q + 1 fresh + variables. The excluded list is at least X and FV M, and then ‘vs’. + But since P is a closed term, these fresh variables seem irrelevant... + *) + >> qabbrev_tac ‘P = permutator d’ + >> ‘closed P’ by rw [Abbr ‘P’, closed_permutator] + >> qabbrev_tac ‘p2 = [[P/y]]’ + >> ‘Boehm_transform p2’ by rw [Abbr ‘p2’] + >> ‘apply p2 M1 = P @* MAP [P/y] args’ by rw [Abbr ‘p2’, appstar_SUB] + >> qabbrev_tac ‘args' = MAP [P/y] args’ >> Know ‘!i. i < m ==> FV (EL i args') SUBSET FV (EL i args)’ - >- (rw [Abbr ‘m’, Abbr ‘args'’, EL_MAP] \\ - qabbrev_tac ‘N = EL i args’ \\ - simp [FV_SUB] \\ - ‘FV P = {}’ by rw [Abbr ‘P’, FV_permutator] \\ - simp [] \\ - Cases_on ‘v IN FV N’ >> SET_TAC []) + >- (rw [Abbr ‘args'’, EL_MAP, FV_SUB] \\ + fs [closed_def]) >> DISCH_TAC - >> Q.PAT_X_ASSUM ‘ltree_lookup (BT' X M r) (h::p) <> NONE’ - (MP_TAC o REWRITE_RULE [BT_def, BT_generator_def, ltree_unfold]) - >> simp [principle_hnf_beta_reduce, hnf_appstar, LMAP_fromList, - ltree_lookup_def, LNTH_fromList] - >> Cases_on ‘h < m’ >> simp [EL_MAP, GSYM BT_def] + >> ‘LENGTH args' = m’ by rw [Abbr ‘args'’, Abbr ‘m’] + (* NOTE: Z contains ‘vs’ in addition to X and FV M *) + >> qabbrev_tac ‘Z = X UNION FV M UNION set vs’ + >> ‘FINITE Z’ by (rw [Abbr ‘Z’] >> rw []) + >> Know ‘solvable (M0 @* MAP VAR vs)’ + >- (MATCH_MP_TAC solvable_appstar' \\ + qexistsl_tac [‘X’, ‘M’, ‘r’, ‘n’] >> simp []) + >> DISCH_TAC + >> Know ‘FV M1 SUBSET Z’ + >- (MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘FV M0 UNION set vs’ \\ + reverse CONJ_TAC + >- (qunabbrev_tac ‘Z’ \\ + Suff ‘FV M0 SUBSET FV M’ >- SET_TAC [] \\ + qunabbrev_tac ‘M0’ \\ + MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) \\ + ‘FV M0 UNION set vs = FV (M0 @* MAP VAR vs)’ by rw [FV_appstar_MAP_VAR] \\ + POP_ORW \\ + qunabbrev_tac ‘M1’ \\ + MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) + >> DISCH_TAC + >> qabbrev_tac ‘z = SUC (string_width Z)’ + >> qabbrev_tac ‘l = alloc r z (z + d - m + 1)’ + >> Know ‘ALL_DISTINCT l /\ LENGTH l = d - m + 1’ + >- rw [Abbr ‘l’, alloc_thm] + >> STRIP_TAC + >> Know ‘DISJOINT (set l) Z’ + >- (rw [Abbr ‘l’, Abbr ‘z’, DISJOINT_ALT', alloc_def, MEM_MAP] \\ + ONCE_REWRITE_TAC [TAUT ‘~P \/ Q \/ ~R <=> P /\ R ==> Q’] \\ + STRIP_TAC \\ + ‘FINITE Z’ by rw [Abbr ‘Z’] \\ + MP_TAC (Q.SPECL [‘x’, ‘Z’] string_width_thm) >> rw []) + >> DISCH_TAC + (* now recover the old definition of Y *) + >> Know ‘DISJOINT (set l) (FV M1)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘Z’ >> art []) + >> ASM_REWRITE_TAC [FV_appstar, FV_thm] + >> DISCH_THEN (STRIP_ASSUME_TAC o (REWRITE_RULE [DISJOINT_UNION'])) + >> Q.PAT_X_ASSUM ‘DISJOINT (set l) {y}’ (* ~MEM y l *) + (STRIP_ASSUME_TAC o (SIMP_RULE (srw_ss()) [DISJOINT_ALT'])) + >> ‘l <> []’ by rw [NOT_NIL_EQ_LENGTH_NOT_0] + >> qabbrev_tac ‘as = FRONT l’ + >> ‘LENGTH as = d - m’ by rw [Abbr ‘as’, LENGTH_FRONT] + >> qabbrev_tac ‘b = LAST l’ + >> Know ‘l = SNOC b as’ + >- (ASM_SIMP_TAC std_ss [Abbr ‘as’, Abbr ‘b’, SNOC_LAST_FRONT]) + >> DISCH_TAC + >> qabbrev_tac ‘p3 = MAP rightctxt (REVERSE (MAP VAR l))’ + >> ‘Boehm_transform p3’ by rw [Abbr ‘p3’, MAP_MAP_o, GSYM MAP_REVERSE] + (* applying permutator_thm *) + >> Know ‘apply p3 (P @* args') == VAR b @* args' @* MAP VAR as’ + >- (simp [Abbr ‘p3’, Abbr ‘P’, rightctxt_thm, MAP_SNOC, + Boehm_apply_MAP_rightctxt'] \\ + REWRITE_TAC [GSYM appstar_APPEND, APPEND_ASSOC] \\ + MATCH_MP_TAC permutator_thm >> rw []) + >> DISCH_TAC + (* pre-final stage *) + >> Q.EXISTS_TAC ‘p3 ++ p2 ++ p1’ + >> CONJ_ASM1_TAC + >- (MATCH_MP_TAC Boehm_transform_APPEND >> art [] \\ + MATCH_MP_TAC Boehm_transform_APPEND >> art []) + >> qabbrev_tac ‘pi = p3 ++ p2 ++ p1’ + >> Know ‘apply pi M == VAR b @* args' @* MAP VAR as’ + >- (MATCH_MP_TAC lameq_TRANS \\ + Q.EXISTS_TAC ‘apply pi M0’ \\ + CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong \\ + CONJ_TAC >- art [] \\ + qunabbrev_tac ‘M0’ \\ + MATCH_MP_TAC lameq_SYM \\ + MATCH_MP_TAC lameq_principle_hnf' >> art []) \\ + SIMP_TAC std_ss [Once Boehm_apply_APPEND, Abbr ‘pi’] \\ + MATCH_MP_TAC lameq_TRANS \\ + Q.EXISTS_TAC ‘apply (p3 ++ p2) M1’ \\ + CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong >> art [] \\ + MATCH_MP_TAC Boehm_transform_APPEND >> art []) \\ + ONCE_REWRITE_TAC [Boehm_apply_APPEND] \\ + MATCH_MP_TAC lameq_TRANS \\ + Q.EXISTS_TAC ‘apply p3 (P @* args')’ >> art [] \\ + MATCH_MP_TAC Boehm_apply_lameq_cong >> rw []) + >> DISCH_TAC + >> CONJ_ASM1_TAC + >- (Suff ‘solvable (VAR b @* args' @* MAP VAR as)’ + >- METIS_TAC [lameq_solvable_cong] \\ + simp [solvable_iff_has_hnf, GSYM appstar_APPEND] \\ + MATCH_MP_TAC hnf_has_hnf >> simp [hnf_appstar]) + (* stage work *) + >> Know ‘principle_hnf (apply pi M) = VAR b @* args' @* MAP VAR as’ + >- (Q.PAT_X_ASSUM ‘Boehm_transform pi’ K_TAC \\ + Q.PAT_X_ASSUM ‘Boehm_transform p1’ K_TAC \\ + Q.PAT_X_ASSUM ‘Boehm_transform p2’ K_TAC \\ + Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC \\ + Q.PAT_X_ASSUM ‘apply p1 M0 == M1’ K_TAC \\ + Q.PAT_X_ASSUM ‘apply p2 M1 = P @* args'’ K_TAC \\ + Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC \\ + (* preparing for principle_hnf_denude_thm *) + Q.PAT_X_ASSUM ‘apply pi M == _’ MP_TAC \\ + simp [Boehm_apply_APPEND, Abbr ‘pi’, Abbr ‘p1’, Abbr ‘p2’, Abbr ‘p3’, + Boehm_apply_MAP_rightctxt'] \\ + Know ‘[P/y] (M @* MAP VAR vs) @* MAP VAR (SNOC b as) = + [P/y] (M @* MAP VAR vs @* MAP VAR (SNOC b as))’ + >- (simp [appstar_SUB] \\ + Suff ‘MAP [P/y] (MAP VAR (SNOC b as)) = MAP VAR (SNOC b as)’ >- Rewr \\ + Q.PAT_X_ASSUM ‘l = SNOC b as’ (ONCE_REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘LENGTH l = d - m + 1’ K_TAC \\ + rw [LIST_EQ_REWRITE, EL_MAP] \\ + MATCH_MP_TAC lemma14b \\ + REWRITE_TAC [FV_thm, IN_SING] \\ + Q.PAT_X_ASSUM ‘~MEM y l’ MP_TAC \\ + rw [MEM_EL] >> METIS_TAC []) >> Rewr' \\ + DISCH_TAC (* [P/y] ... == ... *) \\ + (* applying principle_hnf_permutator *) + Know ‘VAR b @* args' @* MAP VAR as = + principle_hnf ([P/y] (VAR y @* args @* MAP VAR (SNOC b as)))’ + >- (ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ + simp [appstar_SUB, appstar_SNOC, MAP_SNOC] \\ + Know ‘MAP [P/y] (MAP VAR as) = MAP VAR as’ + >- (Q.PAT_X_ASSUM ‘LENGTH as = _’ K_TAC \\ + rw [LIST_EQ_REWRITE, EL_MAP] \\ + MATCH_MP_TAC lemma14b >> rw [] \\ + Q.PAT_X_ASSUM ‘~MEM y (SNOC b as)’ MP_TAC \\ + rw [MEM_EL] >> PROVE_TAC []) >> Rewr' \\ + Know ‘[P/y] (VAR b) = VAR b’ + >- (MATCH_MP_TAC lemma14b >> fs [MEM_SNOC, IN_UNION]) >> Rewr' \\ + simp [Abbr ‘P’, GSYM appstar_APPEND] \\ + MATCH_MP_TAC principle_hnf_permutator >> rw []) >> Rewr' \\ + (* applying principle_hnf_SUB_cong *) + MATCH_MP_TAC principle_hnf_SUB_cong \\ + CONJ_TAC (* has_hnf #1 *) + >- (simp [GSYM solvable_iff_has_hnf, GSYM appstar_APPEND] \\ + ‘M0 == M’ by rw [lameq_principle_hnf', Abbr ‘M0’] \\ + ‘M0 @* (MAP VAR vs ++ MAP VAR (SNOC b as)) == + M @* (MAP VAR vs ++ MAP VAR (SNOC b as))’ by rw [lameq_appstar_cong] \\ + Suff ‘solvable (M0 @* (MAP VAR vs ++ MAP VAR (SNOC b as)))’ + >- PROVE_TAC [lameq_solvable_cong] \\ + NTAC 2 (POP_ASSUM K_TAC) \\ + REWRITE_TAC [appstar_APPEND] \\ + qabbrev_tac ‘M0' = M0 @* MAP VAR vs’ \\ + ‘M0' == M1’ by rw [Abbr ‘M0'’] \\ + ‘M0' @* MAP VAR (SNOC b as) == M1 @* MAP VAR (SNOC b as)’ + by rw [lameq_appstar_cong] \\ + Suff ‘solvable (M1 @* MAP VAR (SNOC b as))’ + >- PROVE_TAC [lameq_solvable_cong] \\ + REWRITE_TAC [solvable_iff_has_hnf] \\ + MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ + CONJ_TAC (* has_hnf #2 *) + >- (REWRITE_TAC [GSYM solvable_iff_has_hnf] \\ + Suff ‘solvable (VAR b @* args' @* MAP VAR as)’ + >- PROVE_TAC [lameq_solvable_cong] \\ + REWRITE_TAC [solvable_iff_has_hnf] \\ + MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ + CONJ_TAC (* has_hnf # 3 *) + >- (simp [appstar_SUB, MAP_SNOC] \\ + Know ‘MAP [P/y] (MAP VAR as) = MAP VAR as’ + >- (Q.PAT_X_ASSUM ‘LENGTH as = _’ K_TAC \\ + rw [LIST_EQ_REWRITE, EL_MAP] \\ + MATCH_MP_TAC lemma14b >> rw [] \\ + Q.PAT_X_ASSUM ‘~MEM y (SNOC b as)’ MP_TAC \\ + rw [MEM_EL] >> PROVE_TAC []) >> Rewr' \\ + Know ‘[P/y] (VAR b) = VAR b’ + >- (MATCH_MP_TAC lemma14b >> fs [MEM_SNOC, IN_UNION]) >> Rewr' \\ + simp [Abbr ‘P’, GSYM appstar_APPEND] \\ + REWRITE_TAC [GSYM solvable_iff_has_hnf] \\ + Know ‘permutator d @* (args' ++ MAP VAR as) @@ VAR b == + VAR b @* (args' ++ MAP VAR as)’ + >- (MATCH_MP_TAC permutator_thm >> rw []) >> DISCH_TAC \\ + Suff ‘solvable (VAR b @* (args' ++ MAP VAR as))’ + >- PROVE_TAC [lameq_solvable_cong] \\ + REWRITE_TAC [solvable_iff_has_hnf] \\ + MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) \\ + (* applying the celebrating principle_hnf_denude_thm + + NOTE: here ‘DISJOINT (set vs) (FV M)’ is required, and this means that + ‘vs’ must exclude ‘FV M’ instead of just ‘FV M0’. + *) + MATCH_MP_TAC principle_hnf_denude_thm >> rw []) >> DISCH_TAC - >> Know ‘subterm_width M (h::p) <= d <=> - m <= d /\ subterm_width (EL h args) p <= d’ - >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ - qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘vs’, ‘M1’] >> simp []) - >> DISCH_THEN (rfs o wrap) - >> qabbrev_tac ‘N = EL h args’ - (* stage work *) - >> reverse (Cases_on ‘y = v’) - >- (Q.PAT_X_ASSUM ‘[P/v] M -h->* _’ MP_TAC \\ - simp [] >> DISCH_TAC \\ - qabbrev_tac ‘M0' = LAMl vs (VAR y @* args')’ \\ - ‘hnf M0'’ by rw [hnf_appstar, Abbr ‘M0'’] \\ - ‘principle_hnf ([P/v] M) = M0'’ by METIS_TAC [principle_hnf_thm'] \\ - Q.PAT_X_ASSUM ‘hnf M0'’ K_TAC \\ - Q.PAT_X_ASSUM ‘M0 = LAMl vs (VAR y @* args)’ (ASSUME_TAC o SYM) \\ - Q.PAT_X_ASSUM ‘M1 = VAR y @* args’ (ASSUME_TAC o SYM) \\ - ‘hnf_children M1 = args’ by rw [hnf_children_hnf] \\ - ‘LAMl_size M0' = n’ by rw [Abbr ‘M0'’, LAMl_size_hnf] \\ - ‘principle_hnf (M0' @* MAP VAR vs) = VAR y @* args'’ - by rw [Abbr ‘M0'’, principle_hnf_beta_reduce, hnf_appstar] \\ - STRONG_CONJ_TAC - >- (simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def, - LNTH_fromList, EL_MAP] \\ - simp [GSYM BT_def, EL_MAP, Abbr ‘args'’] \\ - FIRST_X_ASSUM (MATCH_MP_TAC o cj 1) >> art [] \\ - CONJ_TAC - >- (qunabbrev_tac ‘N’ \\ - MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] \\ - Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ - simp []) \\ - Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ - qunabbrev_tac ‘Y’ \\ - Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ - >- METIS_TAC [SUBSET_DEF] \\ - Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) >> DISCH_TAC \\ - Know ‘subterm_width ([P/v] M) (h::p) <= d <=> - m <= d /\ subterm_width (EL h args') p <= d’ - >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ - qexistsl_tac [‘X’, ‘r’, ‘M0'’, ‘n’, ‘vs’, ‘VAR y @* args'’] \\ - simp [principle_hnf_beta_reduce, ltree_paths_def] \\ - CONJ_TAC - >- (rw [FV_SUB, Abbr ‘P’, FV_permutator] \\ - MATCH_MP_TAC SUBSET_TRANS >> Q.EXISTS_TAC ‘FV M’ >> art [] \\ - SET_TAC []) \\ - fs [Abbr ‘m’, Abbr ‘args'’, Abbr ‘M0'’]) >> Rewr' \\ - simp [Abbr ‘args'’, EL_MAP] \\ - FIRST_X_ASSUM (MATCH_MP_TAC o cj 2) >> art [] \\ - Q.EXISTS_TAC ‘SUC r’ >> art [] \\ + (* applying is_ready_alt *) + >> CONJ_TAC + >- (simp [is_ready_alt, Abbr ‘pi’] \\ + qexistsl_tac [‘b’, ‘args' ++ MAP VAR as’] \\ CONJ_TAC - >- (qunabbrev_tac ‘N’ \\ - MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp [] \\ - Q.PAT_X_ASSUM ‘_ = M0’ (REWRITE_TAC o wrap o SYM) \\ - simp []) \\ - Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC \\ + >- (MP_TAC (Q.SPEC ‘VAR b @* args' @* MAP VAR as’ + (MATCH_MP principle_hnf_thm' + (ASSUME “solvable (apply (p3 ++ p2 ++ p1) M)”))) \\ + simp [appstar_APPEND]) \\ + simp [] (* now two EVERY *) \\ + reverse CONJ_TAC + >- (rw [EVERY_MEM, Abbr ‘b’, Abbr ‘as’, MEM_MAP] >> rw [FV_thm] \\ + Q.PAT_X_ASSUM ‘ALL_DISTINCT l’ MP_TAC \\ + Q.PAT_X_ASSUM ‘l = SNOC (LAST l) (FRONT l)’ (ONCE_REWRITE_TAC o wrap) \\ + rw [ALL_DISTINCT_SNOC] >> PROVE_TAC []) \\ + rw [EVERY_MEM, MEM_MAP] \\ + qabbrev_tac ‘Y = BIGUNION (IMAGE FV (set args))’ \\ + rfs [LIST_TO_SET_SNOC] \\ + Suff ‘FV e SUBSET Y’ >- METIS_TAC [SUBSET_DEF] \\ qunabbrev_tac ‘Y’ \\ - Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ - >- METIS_TAC [SUBSET_DEF] \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘BIGUNION (IMAGE FV (set args'))’ \\ + reverse CONJ_TAC + >- (rw [SUBSET_DEF, IN_BIGUNION_IMAGE, MEM_EL] \\ + Q.EXISTS_TAC ‘EL n args’ \\ + CONJ_TAC >- (Q.EXISTS_TAC ‘n’ >> art []) \\ + POP_ASSUM MP_TAC \\ + Suff ‘FV (EL n args') SUBSET FV (EL n args)’ >- METIS_TAC [SUBSET_DEF] \\ + FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ + rw [SUBSET_DEF, IN_BIGUNION_IMAGE] \\ + Q.EXISTS_TAC ‘e’ >> art []) + (* extra goal *) + >> CONJ_TAC + >- (Q.PAT_X_ASSUM ‘apply pi M == _’ K_TAC \\ + Q.PAT_X_ASSUM ‘principle_hnf (apply pi M) = _’ K_TAC \\ + Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC \\ + rpt (Q.PAT_X_ASSUM ‘Boehm_transform _’ K_TAC) \\ + Q.PAT_X_ASSUM ‘solvable (apply pi M)’ K_TAC \\ + simp [Boehm_apply_APPEND, Abbr ‘pi’, Abbr ‘p1’, Abbr ‘p2’, Abbr ‘p3’, + Boehm_apply_MAP_rightctxt'] \\ + POP_ASSUM (ONCE_REWRITE_TAC o wrap o SYM) \\ + reverse CONJ_TAC + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW r’ \\ + rw [Abbr ‘l’, alloc_SUBSET_ROW] \\ + Suff ‘ROW r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [ROW_SUBSET_RANK]) \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘FV (M @* MAP VAR vs)’ \\ + CONJ_TAC >- (MATCH_MP_TAC FV_SUB_SUBSET >> art []) \\ + simp [FV_appstar] \\ + reverse CONJ_TAC + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW r’ \\ + rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ + Suff ‘ROW r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [ROW_SUBSET_RANK]) \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ >> art [] \\ Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ rw [RANK_MONO]) - (* stage work *) - >> POP_ASSUM (rfs o wrap o SYM) (* [P/y] from now on *) - >> Q.PAT_X_ASSUM ‘[P/y] M -h->* _’ MP_TAC + (* stage work, there's the textbook choice of y and P *) + >> qexistsl_tac [‘y’, ‘P’] >> art [] + >> NTAC 2 STRIP_TAC (* push ‘q <<= p’ to assumptions *) + (* RHS rewriting from M to M0 *) + >> Know ‘subterm X M0 q r = subterm X M q r’ + >- (qunabbrev_tac ‘M0’ \\ + MATCH_MP_TAC subterm_of_principle_hnf >> art []) + >> DISCH_THEN (ONCE_REWRITE_TAC o wrap o SYM) + (* LHS rewriting from M to M0 *) + >> Know ‘subterm X (apply pi M) q r = + subterm X (VAR b @* args' @* MAP VAR as) q r’ + >- (Q.PAT_X_ASSUM ‘_ = VAR b @* args' @* MAP VAR as’ + (ONCE_REWRITE_TAC o wrap o SYM) \\ + qabbrev_tac ‘t = apply pi M’ \\ + ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ + MATCH_MP_TAC subterm_of_principle_hnf >> art []) + >> Rewr' + (* stage cleanups *) + >> Q.PAT_X_ASSUM ‘solvable (apply pi M)’ K_TAC + >> Q.PAT_X_ASSUM ‘principle_hnf (apply pi M) = _’ K_TAC + >> Q.PAT_X_ASSUM ‘apply pi M == _’ K_TAC + >> Q.PAT_X_ASSUM ‘Boehm_transform pi’ K_TAC + (* stage work, now ‘M’ is eliminated from both sides! *) + >> Cases_on ‘q’ >- FULL_SIMP_TAC std_ss [] (* this asserts q = h::t *) + >> Know ‘h < m’ + >- (Cases_on ‘p’ >> fs [] \\ + Q.PAT_X_ASSUM ‘h = h'’ (fs o wrap o SYM) \\ + Know ‘subterm X M (h::t) r <> NONE’ + >- (FIRST_X_ASSUM MATCH_MP_TAC >> rw []) \\ + CONV_TAC (UNBETA_CONV “subterm X M (h::t) r”) \\ + qmatch_abbrev_tac ‘f _’ \\ + RW_TAC bool_ss [subterm_of_solvables] \\ + simp [Abbr ‘f’]) + >> DISCH_TAC + (* applying subterm_of_absfree_hnf *) + >> MP_TAC (Q.SPECL [‘X’, ‘VAR b @* args' @* MAP VAR as’, ‘h’, ‘t’, ‘r’] + subterm_of_absfree_hnf) + >> simp [hnf_appstar, GSYM appstar_APPEND, hnf_children_appstar] + >> DISCH_THEN K_TAC (* already used *) + (* eliminating ‘MAP VAR as’ *) + >> Know ‘EL h (args' ++ MAP VAR as) = EL h args'’ + >- (MATCH_MP_TAC EL_APPEND1 >> rw []) + >> Rewr' + (* eliminating ‘vs’ + + NOTE: ‘subterm Y’ changed to ‘subterm Z’ at next level. There's no + flexibility here on the choice of excluded variabes. + *) + >> Know ‘subterm X (LAMl vs (VAR y @* args)) (h::t) r = + subterm X (EL h args) t (SUC r)’ + >- (MP_TAC (Q.SPECL [‘X’, ‘LAMl vs (VAR y @* args)’, ‘h’, ‘t’, ‘r’] + subterm_of_hnf) \\ + simp [hnf_LAMl, hnf_appstar] \\ + DISCH_THEN K_TAC (* already used *) \\ + Q.PAT_X_ASSUM ‘M0 = LAMl vs (VAR y @* args)’ (REWRITE_TAC o wrap o SYM) \\ + simp [hnf_children_hnf]) + >> Rewr' + (* Now: subterm' Z (EL h args') t == [P/y] (subterm' Z (EL h args) t) + + First of all, those assumptions about p1,p2,p3 are no more needed. + *) + >> qunabbrev_tac ‘pi’ + >> Q.PAT_X_ASSUM ‘Boehm_transform p1’ K_TAC + >> Q.PAT_X_ASSUM ‘apply p1 M0 == M1’ K_TAC + >> qunabbrev_tac ‘p1’ + >> Q.PAT_X_ASSUM ‘Boehm_transform p2’ K_TAC + >> Q.PAT_X_ASSUM ‘apply p2 M1 = P @* args'’ K_TAC + >> qunabbrev_tac ‘p2’ + >> Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC + >> Q.PAT_X_ASSUM ‘apply p3 (P @* args') == _’ K_TAC + >> qunabbrev_tac ‘p3’ + >> Q.PAT_X_ASSUM ‘h::t <> []’ K_TAC + >> qabbrev_tac ‘N = EL h args’ + >> qabbrev_tac ‘N' = EL h args'’ + (* eliminating N' *) + >> ‘N' = [P/y] N’ by simp [EL_MAP, Abbr ‘m’, Abbr ‘N’, Abbr ‘N'’, Abbr ‘args'’] + >> POP_ORW + >> qunabbrev_tac ‘N'’ + (* cleanup args' *) + >> Q.PAT_X_ASSUM + ‘!i. i < m ==> + FV (EL i args') SUBSET FV (EL i args)’ K_TAC + >> Q.PAT_X_ASSUM ‘LENGTH args' = m’ K_TAC + >> qunabbrev_tac ‘args'’ + (* cleanup l, as and b *) + >> Q.PAT_X_ASSUM ‘ALL_DISTINCT l’ K_TAC + >> NTAC 2 (Q.PAT_X_ASSUM ‘DISJOINT (set l) _’ K_TAC) + >> Q.PAT_X_ASSUM ‘LENGTH l = q - m + 1’ K_TAC + >> Q.PAT_X_ASSUM ‘l <> []’ K_TAC + >> Q.PAT_X_ASSUM ‘l = SNOC b as’ K_TAC + >> Q.PAT_X_ASSUM ‘~MEM y l’ K_TAC + >> Q.PAT_X_ASSUM ‘LENGTH as = q - m’ K_TAC + >> qunabbrevl_tac [‘l’, ‘as’, ‘b’] + (* applying subterm_headvar_lemma *) + >> Know ‘hnf_headvar M1 IN X UNION RANK (SUC r)’ + >- (MATCH_MP_TAC subterm_headvar_lemma \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘vs’] >> simp []) + >> ASM_SIMP_TAC std_ss [hnf_head_hnf, THE_VAR_thm] + >> DISCH_TAC (* y IN X UNION RANK (SUC r) *) + (* applying subterm_subst_permutator_cong *) + >> MATCH_MP_TAC subterm_subst_permutator_cong' + >> Q.EXISTS_TAC ‘d’ >> simp [Abbr ‘P’] - >> DISCH_TAC (* [permutator d/v] M -h->* ... *) - (* NOTE: New frash variables xs and y will be asserted here: + >> CONJ_TAC + >- (qunabbrev_tac ‘N’ >> MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) + (* stage work *) + >> CONJ_ASM1_TAC (* subterm Z N t <> NONE *) + >- (Q.PAT_X_ASSUM ‘!q. q <<= p ==> subterm X M q r <> NONE’ + (MP_TAC o Q.SPEC ‘h::t’) \\ + Q.PAT_X_ASSUM ‘M0 = _’ K_TAC \\ + simp [subterm_of_solvables]) + (* final goal: subterm_width (EL h args) t <= d *) + >> qunabbrev_tac ‘d’ + (* applying subterm_width_thm *) + >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘p’, ‘r’] subterm_width_thm) + >> simp [] >> DISCH_THEN K_TAC + (* applying subterm_width_thm again *) + >> MP_TAC (Q.SPECL [‘X’, ‘N’, ‘t’, ‘SUC r’] subterm_width_thm) + >> Know ‘FV N SUBSET X UNION RANK (SUC r)’ + >- (qunabbrev_tac ‘N’ \\ + MATCH_MP_TAC subterm_induction_lemma' \\ + qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) + >> DISCH_TAC + >> ‘t IN ltree_paths (BT' X N (SUC r))’ + by PROVE_TAC [subterm_imp_ltree_paths] + >> simp [] >> DISCH_THEN K_TAC + (* applying SUBSET_MAX_SET *) + >> MATCH_MP_TAC SUBSET_MAX_SET + >> CONJ_TAC + >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) + >> CONJ_TAC + >- (MATCH_MP_TAC IMAGE_FINITE >> rw [FINITE_prefix]) + (* final goal *) + >> ‘hnf_children_size M0 = m’ by rw [Abbr ‘m’] + >> Q.PAT_X_ASSUM ‘M0 = _’ K_TAC + >> rw [SUBSET_DEF] (* this asserts q <<= t *) + >> Know ‘h::q <<= p’ + >- (MATCH_MP_TAC IS_PREFIX_TRANS \\ + Q.EXISTS_TAC ‘h::t’ >> simp []) + >> DISCH_TAC + >> Q.EXISTS_TAC ‘h::q’ >> simp [] + >> Know ‘subterm X M (h::q) r <> NONE’ + >- (FIRST_X_ASSUM MATCH_MP_TAC >> simp []) + >> Know ‘subterm X N q (SUC r) <> NONE’ + >- (Cases_on ‘t = []’ >- fs [] \\ + irule (cj 1 subterm_solvable_lemma) >> simp [] \\ + Q.EXISTS_TAC ‘t’ >> art []) + >> DISCH_TAC + >> Q_TAC (UNBETA_TAC [subterm_of_solvables]) ‘subterm' X M (h::q) r’ + >> simp [] +QED - P @* args' -h->* LAMl xs (LAM z (VAR z @* args' @* MAP VAR xs)), +(* Proposition 10.3.7 (i) [1, p.248] (Boehm out lemma) - thus ‘principle_hnf ([P/y] M) = LAMl (vs ++ xs ++ [z]) (VAR z @* ...)’ + NOTE: This time the case ‘p = []’ can be included, but it's a trvial case. - Here LENGTH xs = d - m. Let n' be the LAMl_size of the above hnf. - *) - >> qabbrev_tac ‘n' = n + (d - m) + 1’ - >> Q_TAC (RNEWS_TAC (“zs :string list”, “r :num”, “n' :num”)) ‘X’ - >> MP_TAC (Q.SPECL [‘set zs’, ‘d’, ‘args'’] hreduce_permutator_thm) - >> impl_tac >- rpt (rw []) - >> REWRITE_TAC [DISJOINT_UNION] + NOTE: The original lemma in textbook requires ‘p IN ltree_paths (BT X M)’, + but this seems wrong, as ‘subterm X M p’ may not be defined if only ‘p’ is + valid path (i.e. the subterm could be a bottom (\bot) as the result of un- + solvable terms). + + NOTE: Can we enhance this lemma by using ‘-h->*’ instead of ‘==’? + + NOTE: Use subterm_imp_ltree_paths to prove ‘p IN ltree_paths (BT X M)’. + *) +Theorem Boehm_out_lemma : + !X p M r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + subterm X M p r <> NONE ==> + ?pi. Boehm_transform pi /\ + ?ss. apply pi M == subterm' X M p r ISUB ss +Proof + Q.X_GEN_TAC ‘X’ + >> Induct_on ‘p’ + >- (rw [] \\ + Q.EXISTS_TAC ‘[]’ >> rw [] \\ + Q.EXISTS_TAC ‘[]’ >> rw []) + >> rpt STRIP_TAC + >> rename1 ‘subterm X M (h::t) r <> NONE’ + >> qabbrev_tac ‘p = h::t’ (* head and tail *) + >> ‘p <> []’ by rw [Abbr ‘p’] + >> ‘(!q. q <<= p ==> subterm X M q r <> NONE) /\ + !q. q <<= FRONT p ==> solvable (subterm' X M q r)’ + by METIS_TAC [subterm_solvable_lemma] + >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘h::t’, ‘r’] Boehm_transform_exists_lemma) + >> simp [] + >> DISCH_THEN (Q.X_CHOOSE_THEN ‘p0’ MP_TAC) + >> RW_TAC std_ss [] (* push p0 properties to assumptions *) + >> POP_ASSUM (MP_TAC o (Q.SPEC ‘p’)) (* put q = p *) + >> rw [] + >> qabbrev_tac ‘M' = apply p0 M’ + >> Q.PAT_X_ASSUM ‘is_ready M'’ (MP_TAC o REWRITE_RULE [is_ready_alt]) >> STRIP_TAC - >> rename1 ‘ALL_DISTINCT (SNOC z xs)’ (* y' -> z *) - >> qabbrev_tac ‘P = permutator d’ - >> Know ‘LAMl vs (P @* args') -h->* - LAMl vs (LAMl xs (LAM z (VAR z @* args' @* MAP VAR xs)))’ - >- (rw [hreduce_LAMl]) - >> qmatch_abbrev_tac ‘LAMl vs (P @* args') -h->* t ==> _’ - >> DISCH_TAC - >> ‘hnf t’ by rw [Abbr ‘t’, hnf_appstar, hnf_LAMl] - >> ‘[P/y] M -h->* t’ by PROVE_TAC [hreduce_TRANS] - >> ‘principle_hnf ([P/y] M) = t’ by METIS_TAC [principle_hnf_thm'] - (* cleanup *) - >> Q.PAT_X_ASSUM ‘P @* args' -h->* _’ K_TAC - >> Q.PAT_X_ASSUM ‘LAMl vs _ -h->* t’ K_TAC - >> Q.PAT_X_ASSUM ‘[P/y] M -h->* LAMl vs _’ K_TAC - >> Q.PAT_X_ASSUM ‘[P/y] M -h->* t’ K_TAC - >> Q.PAT_X_ASSUM ‘hnf t’ K_TAC - >> qunabbrev_tac ‘t’ + >> ‘principle_hnf M' = VAR y @* Ns’ by rw [principle_hnf_thm', hnf_appstar] (* stage work *) - >> POP_ASSUM MP_TAC - >> ‘!t. LAMl vs (LAMl xs (LAM z t)) = LAMl (vs ++ (SNOC z xs)) t’ - by rw [LAMl_APPEND] - >> POP_ORW - >> REWRITE_TAC [GSYM appstar_APPEND] - >> qabbrev_tac ‘args'' = args' ++ MAP VAR xs’ - >> ‘LENGTH args'' = d’ by rw [Abbr ‘args''’] - >> qabbrev_tac ‘xs' = SNOC z xs’ - >> ‘LENGTH xs' = d - m + 1’ by rw [Abbr ‘xs'’] - >> qabbrev_tac ‘vs' = vs ++ xs'’ - >> DISCH_TAC (* principle_hnf ([P/y] M) = ... *) - >> Know ‘LENGTH vs' = n'’ - >- (qunabbrevl_tac [‘n’, ‘n'’, ‘vs'’, ‘xs'’] \\ - Q.PAT_X_ASSUM ‘M0 = _’ (REWRITE_TAC o wrap) \\ - simp [LAMl_size_LAMl]) - >> DISCH_TAC - (* applying NEWS_prefix *) - >> Know ‘vs <<= zs’ - >- (qunabbrevl_tac [‘vs’, ‘zs’] \\ - MATCH_MP_TAC RNEWS_prefix >> rw [Abbr ‘n'’]) - >> DISCH_THEN (STRIP_ASSUME_TAC o (REWRITE_RULE [IS_PREFIX_APPEND])) - >> rename1 ‘zs = vs ++ ys'’ (* l -> ys' *) - >> Know ‘LENGTH ys' = d - m + 1’ - >- (Know ‘LENGTH zs = LENGTH (vs ++ ys')’ >- POP_ASSUM (REWRITE_TAC o wrap) \\ - simp [Abbr ‘n'’]) - >> DISCH_TAC - >> Know ‘!N. MEM N args' ==> DISJOINT (FV N) (set ys')’ - >- (Q.X_GEN_TAC ‘x’ \\ - simp [MEM_EL] \\ - DISCH_THEN (Q.X_CHOOSE_THEN ‘i’ STRIP_ASSUME_TAC) \\ - POP_ORW \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘FV (EL i args)’ >> simp [] \\ - (* applying FV_subterm_lemma *) - Know ‘FV (EL i args) SUBSET FV M UNION set vs’ - >- (MATCH_MP_TAC FV_subterm_lemma \\ - qexistsl_tac [‘X’, ‘r’, ‘M0’, ‘n’, ‘m’, ‘M1’] >> simp []) \\ - DISCH_TAC \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘FV M UNION set vs’ >> art [] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘X UNION RANK r UNION set vs’ \\ - reverse CONJ_TAC >- ASM_SET_TAC [] \\ - simp [DISJOINT_ALT'] \\ - rpt CONJ_TAC - >- (NTAC 2 STRIP_TAC \\ - Q.PAT_X_ASSUM ‘DISJOINT (set zs) X’ MP_TAC \\ - rw [DISJOINT_ALT]) - >- (NTAC 2 STRIP_TAC \\ - MP_TAC (Q.SPECL [‘r’, ‘n'’, ‘X’] DISJOINT_RANK_RNEWS') \\ - rw [DISJOINT_ALT', IN_UNION]) \\ - Q.PAT_X_ASSUM ‘ALL_DISTINCT zs’ MP_TAC \\ - rw [ALL_DISTINCT_APPEND] >> METIS_TAC []) - >> DISCH_TAC - >> qabbrev_tac ‘t = VAR z @* args''’ - >> ‘hnf t’ by rw [Abbr ‘t’, hnf_appstar] - >> qabbrev_tac ‘M0' = LAMl vs' t’ - >> ‘LAMl_size M0' = n'’ by rw [Abbr ‘M0'’, Abbr ‘t’] - >> qabbrev_tac ‘M1' = principle_hnf (M0' @* MAP VAR zs)’ - >> Know ‘M1' = tpm (ZIP (xs',ys')) t’ - >- (simp [Abbr ‘M0'’, Abbr ‘M1'’, Abbr ‘vs'’, GSYM APPEND_ASSOC, appstar_APPEND, - LAMl_APPEND] \\ - Know ‘principle_hnf (LAMl vs (LAMl xs' t) @* MAP VAR vs @* MAP VAR ys') = - principle_hnf (LAMl xs' t @* MAP VAR ys')’ - >- (MATCH_MP_TAC principle_hnf_hreduce \\ - simp [hreduce_BETA_extended]) >> Rewr' \\ - MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ - CONJ_TAC >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT zs’ MP_TAC \\ - Q.PAT_X_ASSUM ‘zs = vs ++ ys'’ (ONCE_REWRITE_TAC o wrap) \\ - simp [ALL_DISTINCT_APPEND]) \\ - CONJ_ASM1_TAC - >- (ONCE_REWRITE_TAC [DISJOINT_SYM] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set zs’ >> art [] \\ - rw [LIST_TO_SET_APPEND]) \\ - simp [Abbr ‘t’, Abbr ‘args''’, FV_appstar] \\ - CONJ_TAC (* ~MEM z ys' *) - >- (POP_ASSUM MP_TAC \\ - rw [DISJOINT_ALT, Abbr ‘xs'’]) \\ - reverse CONJ_TAC (* MAP VAR xs *) - >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs') (set ys')’ MP_TAC \\ - rw [Abbr ‘xs'’, DISJOINT_ALT]) \\ - Q.X_GEN_TAC ‘s’ >> simp [MEM_MAP] \\ - rpt STRIP_TAC \\ - Q.PAT_X_ASSUM ‘s = FV x’ (REWRITE_TAC o wrap) \\ - FIRST_X_ASSUM MATCH_MP_TAC >> art []) + >> qunabbrev_tac ‘p’ + >> Know ‘h < LENGTH Ns’ + >- (Q.PAT_X_ASSUM ‘subterm X M' (h::t) r <> NONE’ MP_TAC \\ + RW_TAC std_ss [subterm_of_solvables] >> fs []) >> DISCH_TAC - >> Q.PAT_X_ASSUM ‘zs = vs ++ ys'’ (ASSUME_TAC o SYM) + >> qabbrev_tac ‘m = LENGTH Ns’ + >> qabbrev_tac ‘N = EL h Ns’ (* stage work *) - >> STRONG_CONJ_TAC - >- (simp [BT_def, BT_generator_def, Once ltree_unfold, ltree_lookup_def, - LNTH_fromList, EL_MAP] \\ - REWRITE_TAC [GSYM BT_def] \\ - simp [hnf_children_tpm, EL_MAP] \\ - simp [Abbr ‘t’, hnf_children_hnf] \\ - Know ‘h < LENGTH args''’ - >- (Q_TAC (TRANS_TAC LESS_LESS_EQ_TRANS) ‘m’ >> art [] \\ - rw [Abbr ‘args''’]) >> Rewr \\ - ‘EL h args'' = EL h args'’ by rw [Abbr ‘args''’, EL_APPEND1] \\ - POP_ORW \\ - Know ‘tpm (ZIP (xs',ys')) (EL h args') = EL h args'’ - >- (MATCH_MP_TAC tpm_unchanged >> simp [MAP_ZIP] \\ - ONCE_REWRITE_TAC [DISJOINT_SYM] \\ - CONJ_TAC \\ (* 2 subgoals, same tactics *) - FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) >> Rewr' \\ - simp [Abbr ‘args'’, EL_MAP] \\ - FIRST_X_ASSUM (MATCH_MP_TAC o cj 1) >> art [] \\ - qunabbrev_tac ‘Y’ \\ - reverse CONJ_TAC - >- (Q.PAT_X_ASSUM ‘y IN X UNION RANK r’ MP_TAC \\ - Suff ‘X UNION RANK r SUBSET X UNION RANK (SUC r)’ - >- (REWRITE_TAC [SUBSET_DEF] \\ - DISCH_THEN (REWRITE_TAC o wrap)) \\ - Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) \\ - qunabbrev_tac ‘N’ \\ - MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) - >> DISCH_TAC - (* extra goal for subterm_width *) - >> qabbrev_tac ‘pi = ZIP (xs',ys')’ - >> Q.PAT_X_ASSUM ‘hnf t’ K_TAC - >> rfs [Abbr ‘t’, tpm_appstar] - >> qabbrev_tac ‘Ms = listpm term_pmact pi args''’ - >> Know ‘subterm_width ([P/y] M) (h::p) <= d <=> - d <= d /\ subterm_width (EL h Ms) p <= d’ - >- (MATCH_MP_TAC subterm_width_induction_lemma' \\ - qexistsl_tac [‘X’, ‘r’, ‘M0'’, ‘n'’, ‘zs’, ‘M1'’] \\ - simp [principle_hnf_beta_reduce, ltree_paths_def] \\ - CONJ_TAC - >- (rw [FV_SUB, Abbr ‘P’, FV_permutator] \\ - MATCH_MP_TAC SUBSET_TRANS >> Q.EXISTS_TAC ‘FV M’ >> art [] \\ - SET_TAC []) \\ - simp [Abbr ‘M0'’]) - >> Rewr' - >> simp [Abbr ‘Ms’, EL_MAP] - >> ‘EL h args'' = EL h args'’ by rw [Abbr ‘args''’, EL_APPEND1] - >> POP_ORW - >> qunabbrev_tac ‘pi’ - >> Know ‘tpm (ZIP (xs',ys')) (EL h args') = EL h args'’ - >- (MATCH_MP_TAC tpm_unchanged >> simp [MAP_ZIP] \\ - ONCE_REWRITE_TAC [DISJOINT_SYM] \\ - CONJ_TAC \\ (* 2 subgoals, same tactics *) - FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) - >> Rewr' - >> simp [Abbr ‘args'’, EL_MAP] - >> FIRST_X_ASSUM (MATCH_MP_TAC o cj 2) >> art [] - >> Q.EXISTS_TAC ‘SUC r’ >> art [] + >> Know ‘subterm X N t (SUC r) <> NONE /\ + subterm' X M' (h::t) r = subterm' X N t (SUC r)’ + >- (Q.PAT_X_ASSUM ‘subterm' X M' (h::t) r = + [P/v] (subterm' X M (h::t) r)’ K_TAC \\ + Q.PAT_X_ASSUM ‘subterm X M' (h::t) r <> NONE’ MP_TAC \\ + rw [subterm_of_solvables, Abbr ‘N’]) + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘subterm' X M' (h::t) r = subterm' X N t (SUC r)’ (fs o wrap) + >> T_TAC + (* stage work, now define a selector *) + >> qabbrev_tac ‘U = selector h m’ + >> qabbrev_tac ‘p1 = [[U/y]]’ + >> ‘Boehm_transform p1’ by rw [Abbr ‘p1’] + >> qabbrev_tac ‘p10 = p1 ++ p0’ + >> ‘Boehm_transform p10’ by rw [Abbr ‘p10’, Boehm_transform_APPEND] + (* applying properties of selector (U) *) + >> Know ‘apply p10 M == N’ + >- (rw [Abbr ‘p10’, Boehm_apply_APPEND] \\ + MATCH_MP_TAC lameq_TRANS \\ + Q.EXISTS_TAC ‘apply p1 (principle_hnf M')’ \\ + CONJ_TAC >- (MATCH_MP_TAC Boehm_apply_lameq_cong \\ + CONJ_TAC >- art [] \\ + MATCH_MP_TAC lameq_SYM \\ + MATCH_MP_TAC lameq_principle_hnf' >> art []) \\ + rw [Abbr ‘p1’, appstar_SUB] \\ + Know ‘MAP [U/y] Ns = Ns’ + >- (rw [LIST_EQ_REWRITE, EL_MAP] \\ + MATCH_MP_TAC lemma14b \\ + Q.PAT_X_ASSUM ‘EVERY (\e. y # e) Ns’ MP_TAC \\ + rw [EVERY_MEM, MEM_EL] \\ + POP_ASSUM MATCH_MP_TAC >> rename1 ‘i < m’ \\ + Q.EXISTS_TAC ‘i’ >> art []) >> Rewr' \\ + qunabbrevl_tac [‘U’, ‘N’] \\ + MATCH_MP_TAC selector_thm >> rw [Abbr ‘m’]) + >> DISCH_TAC + (* stage work, now using IH *) + >> Q.PAT_X_ASSUM ‘!M r. _’ (MP_TAC o (Q.SPECL [‘N’, ‘SUC r’])) >> simp [] + >> Know ‘FV N SUBSET X UNION RANK (SUC r)’ + >- (qunabbrevl_tac [‘N’, ‘M'’] \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘FV (apply p0 M)’ >> art [] \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘FV (VAR y @* Ns)’ \\ + reverse CONJ_TAC >- (MATCH_MP_TAC hreduce_FV_SUBSET >> art []) \\ + rw [SUBSET_DEF, FV_appstar, IN_UNION] \\ + DISJ2_TAC \\ + Q.EXISTS_TAC ‘FV (EL h Ns)’ >> art [] \\ + Q.EXISTS_TAC ‘EL h Ns’ >> rw [EL_MEM]) + >> RW_TAC std_ss [] + >> rename1 ‘apply p2 N == _ ISUB ss'’ + >> qabbrev_tac ‘N' = subterm' X M (h::t) r’ + (* final stage *) + >> Q.EXISTS_TAC ‘p2 ++ p10’ >> CONJ_TAC - >- (qunabbrev_tac ‘N’ \\ - MATCH_MP_TAC subterm_induction_lemma' \\ - qexistsl_tac [‘M’, ‘M0’, ‘n’, ‘m’, ‘vs’, ‘M1’] >> simp []) - >> Q.PAT_X_ASSUM ‘v IN Y’ MP_TAC - >> qunabbrev_tac ‘Y’ - >> Suff ‘X UNION RANK r SUBSET X UNION (RANK (SUC r))’ - >- METIS_TAC [SUBSET_DEF] - >> Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] - >> rw [RANK_MONO] + >- (MATCH_MP_TAC Boehm_transform_APPEND >> art []) + >> Q.EXISTS_TAC ‘[(P,v)] ++ ss'’ + >> MATCH_MP_TAC lameq_TRANS + >> Q.EXISTS_TAC ‘apply p2 N’ + >> simp [ISUB_def] + >> rw [Boehm_apply_APPEND] + >> MATCH_MP_TAC Boehm_apply_lameq_cong >> art [] QED -Theorem BT_ltree_paths_isub_cong : - !ys p X M r P d ss. +(*---------------------------------------------------------------------------* + * Faithfulness and agreements of terms + *---------------------------------------------------------------------------*) + +(* Definition 10.2.21 (i) [1, p.238] + + NOTE: For ‘y1 = y2’ to be meaningful, here we assumed that vs1 and vs2 + share the same prefix, i.e. either vs1 <<= vs2 or vs2 <<= vs1. In reality, + we have ‘vs1 = RNEWS r n1 X /\ vs2 = RNEWS r n2 X’ for some X and r. + *) +Definition head_equivalent_def : + head_equivalent ((a1,m1) :BT_node # num option) + ((a2,m2) :BT_node # num option) = + case (a1,a2) of + (SOME (vs1,y1),SOME (vs2,y2)) => + y1 = y2 /\ LENGTH vs1 + THE m2 = LENGTH vs2 + THE m1 + | (SOME _,NONE) => F + | (NONE,SOME _) => F + | (NONE,NONE) => T +End + +Theorem head_equivalent_refl[simp] : + head_equivalent A A +Proof + Cases_on ‘A’ >> rw [head_equivalent_def] + >> Cases_on ‘q’ >> rw [] + >> Cases_on ‘x’ >> rw [] +QED + +Theorem head_equivalent_sym : + !A B. head_equivalent A B ==> head_equivalent B A +Proof + qx_genl_tac [‘A’, ‘B’] + >> Cases_on ‘A’ >> Cases_on ‘B’ >> simp [head_equivalent_def] + >> Cases_on ‘q’ >> Cases_on ‘q'’ >> simp [] + >> Cases_on ‘x’ >> Cases_on ‘x'’ >> simp [] +QED + +Theorem head_equivalent_comm : + !A B. head_equivalent A B <=> head_equivalent B A +Proof + rpt GEN_TAC + >> EQ_TAC >> rw [head_equivalent_sym] +QED + +(* Definition 10.2.21 (ii) [1, p.238] *) +Overload ltree_equiv = “OPTREL head_equivalent” + +Theorem ltree_equiv_refl[simp] : + ltree_equiv A A +Proof + MATCH_MP_TAC OPTREL_refl >> rw [] +QED + +Theorem ltree_equiv_sym : + !A B. ltree_equiv A B ==> ltree_equiv B A +Proof + rpt GEN_TAC + >> Cases_on ‘A’ >> Cases_on ‘B’ >> rw [OPTREL_THM] + >> rw [Once head_equivalent_comm] +QED + +Theorem ltree_equiv_comm : + !A B. ltree_equiv A B <=> ltree_equiv B A +Proof + rpt STRIP_TAC + >> EQ_TAC >> rw [ltree_equiv_sym] +QED + +Theorem ltree_equiv_some_bot_imp : + !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ + ltree_equiv (SOME bot) (ltree_el (BT' X M r) p) ==> + ltree_el (BT' X M r) p = SOME bot +Proof + rw [OPTREL_def] + >> Cases_on ‘y0’ >> fs [head_equivalent_def] + >> Cases_on ‘q’ >> fs [] + >> METIS_TAC [BT_ltree_el_eq_some_none] +QED + +(* |- !X M p r. FINITE X /\ FV M SUBSET X UNION RANK r /\ - P = permutator d /\ - (!y. MEM y ys ==> y IN X UNION RANK r) /\ ss = MAP (\y. (P,y)) ys /\ - p IN ltree_paths (BT' X M r) /\ - subterm_width M p <= d - ==> p IN ltree_paths (BT' X (M ISUB ss) r) /\ - subterm_width (M ISUB ss) p <= d + ltree_equiv (ltree_el (BT' X M r) p) (SOME bot) ==> + ltree_el (BT' X M r) p = SOME bot + *) +Theorem ltree_equiv_some_bot_imp' = + ONCE_REWRITE_RULE [ltree_equiv_comm] ltree_equiv_some_bot_imp + +(* Definition 10.2.32 (v) [1, p.245] + + NOTE: It's assumed that ‘X SUBSET FV M UNION FV N’ in applications. + *) +Definition subtree_equiv_def : + subtree_equiv X M N p r = + ltree_equiv (ltree_el (BT' X M r) p) (ltree_el (BT' X N r) p) +End + +Theorem subtree_equiv_refl[simp] : + subtree_equiv X M M p r Proof - Induct_on ‘ys’ >- rw [] - >> rpt GEN_TAC - >> STRIP_TAC - >> Q.PAT_X_ASSUM ‘ss = _’ (REWRITE_TAC o wrap) - >> simp [] - >> Q.PAT_X_ASSUM ‘P = permutator d’ K_TAC - >> qabbrev_tac ‘P = permutator d’ - >> qabbrev_tac ‘N = [P/h] M’ - >> qabbrev_tac ‘ss = MAP (\y. (P,y)) ys’ - >> MP_TAC (Q.SPECL [‘X’, ‘P’, ‘d’, ‘h’, ‘p’, ‘M’, ‘r’] BT_ltree_paths_subst_cong) - >> simp [] - >> STRIP_TAC - >> FIRST_X_ASSUM MATCH_MP_TAC - >> Q.EXISTS_TAC ‘P’ >> simp [] - >> Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV M’ >> art [] - >> rw [FV_SUB, Abbr ‘N’] - >> rw [Abbr ‘P’, FV_permutator] + rw [subtree_equiv_def] +QED + +Theorem subtree_equiv_comm : + !X M N p r. subtree_equiv X M N p r <=> subtree_equiv X N M p r +Proof + rw [subtree_equiv_def, Once ltree_equiv_comm] +QED + +(* NOTE: This is the explicit form of the Boehm transform constructed in the + next lemma. It assumes (at least): + + 1) FINITE X + 2) BIGUNION (IMAGE FV (set Ms)) SUBSET X UNION RANK r (0 < r) + 3) EVERY solvable Ms + *) +Definition Boehm_construction_def : + Boehm_construction X (Ms :term list) p = + let n_max = MAX_LIST (MAP (\e. subterm_length e p) Ms); + d_max = MAX_LIST (MAP (\e. subterm_width e p) Ms) + n_max; + k = LENGTH Ms; + Y = BIGUNION (IMAGE FV (set Ms)); + vs0 = NEWS (n_max + SUC d_max + k) (X UNION Y); + vs = TAKE n_max vs0; + xs = DROP n_max vs0; + M i = EL i Ms; + M0 i = principle_hnf (M i); + M1 i = principle_hnf (M0 i @* MAP VAR vs); + y i = hnf_headvar (M1 i); + P i = permutator (d_max + i); + p1 = MAP rightctxt (REVERSE (MAP VAR vs)); + p2 = REVERSE (GENLIST (\i. [P i/y i]) k); + p3 = MAP rightctxt (REVERSE (MAP VAR xs)) + in + p3 ++ p2 ++ p1 +End + +Theorem Boehm_construction_transform : + !X Ms p. Boehm_transform (Boehm_construction X Ms p) +Proof + RW_TAC std_ss [Boehm_construction_def] + >> MATCH_MP_TAC Boehm_transform_APPEND + >> reverse CONJ_TAC + >- rw [Abbr ‘p1’, MAP_MAP_o, GSYM MAP_REVERSE] + >> MATCH_MP_TAC Boehm_transform_APPEND + >> CONJ_TAC + >- rw [Abbr ‘p3’, MAP_MAP_o, GSYM MAP_REVERSE] + >> rw [Boehm_transform_def, Abbr ‘p2’, EVERY_GENLIST] QED -(* Lemma 10.3.11 [1. p.251] (was: agree_upto_lemma) +(* This HUGE theorem is an improved version of Lemma 10.3.11 [1. p.251], to be + proved later in ‘lameta_completeTheory’ as [agree_upto_lemma]. NOTE: ‘p <> []’ must be added, otherwise each N in Ns cannot be "is_ready". @@ -5815,35 +5757,38 @@ QED on the construction of Boehm transform ‘pi’. NOTE: This is the LAST theorem of the current theory, because this proof is - so long. Further results (plus users of this lemma) are to be found in next - completenessTheory, e.g. the actual [agree_upto_lemma] - - NOTE: The original 3rd part “subtree_equiv X M N <=> subtree_equiv X (apply - pi M) (apply pi N)” is unprovable. The present form is the provable subset. + so long. Further results (plus users of this lemma) are to be found in the + next lameta_completeTheory. *) -Theorem subterm_equiv_lemma : +Theorem subtree_equiv_lemma_explicit : !X Ms p r. FINITE X /\ p <> [] /\ 0 < r /\ Ms <> [] /\ BIGUNION (IMAGE FV (set Ms)) SUBSET X UNION RANK r /\ EVERY (\M. subterm X M p r <> NONE) Ms ==> - ?pi. Boehm_transform pi /\ EVERY is_ready' (MAP (apply pi) Ms) /\ - (!M. MEM M Ms ==> p IN ltree_paths (BT' X (apply pi M) r)) /\ - (!M N q. MEM M Ms /\ MEM N Ms /\ q <<= p /\ - subtree_equiv X M N q r ==> - subtree_equiv X (apply pi M) (apply pi N) q r) /\ - !M N. MEM M Ms /\ MEM N Ms /\ - subtree_equiv X (apply pi M) (apply pi N) p r ==> - subtree_equiv' X M N p r + let pi = Boehm_construction X Ms p in + Boehm_transform pi /\ + EVERY is_ready (MAP (apply pi) Ms) /\ + EVERY (\M. p IN ltree_paths (BT' X M r)) (MAP (apply pi) Ms) /\ + !M N q. MEM M Ms /\ MEM N Ms /\ q <<= p ==> + (subtree_equiv X M N q r <=> + subtree_equiv X (apply pi M) (apply pi N) q r) Proof - rpt STRIP_TAC + rpt GEN_TAC >> simp [LET_DEF] + >> STRIP_TAC + >> qabbrev_tac ‘Y = BIGUNION (IMAGE FV (set Ms))’ + >> ‘FINITE Y’ by (rw [Abbr ‘Y’] >> rw []) + >> qabbrev_tac ‘pi' = Boehm_construction X Ms p’ + >> CONJ_ASM1_TAC + >- rw [Abbr ‘pi'’, Boehm_construction_transform] + (* original steps *) + >> Q.PAT_X_ASSUM ‘EVERY _ Ms’ MP_TAC + >> DISCH_THEN (STRIP_ASSUME_TAC o (REWRITE_RULE [EVERY_EL])) >> qabbrev_tac ‘k = LENGTH Ms’ - >> Q.PAT_X_ASSUM ‘EVERY P Ms’ MP_TAC - >> rw [EVERY_EL] >> qabbrev_tac ‘M = \i. EL i Ms’ >> fs [] >> Know ‘!i. i < k ==> FV (M i) SUBSET X UNION RANK r’ >- (rpt STRIP_TAC \\ - Q.PAT_X_ASSUM ‘_ SUBSET X UNION RANK r’ MP_TAC \\ - rw [SUBSET_DEF, IN_BIGUNION_IMAGE] \\ + Q.PAT_X_ASSUM ‘Y SUBSET X UNION RANK r’ MP_TAC \\ + rw [Abbr ‘Y’, SUBSET_DEF, IN_BIGUNION_IMAGE] \\ FIRST_X_ASSUM MATCH_MP_TAC >> Q.EXISTS_TAC ‘M i’ >> art [] \\ rw [Abbr ‘M’, EL_MEM]) >> DISCH_TAC @@ -5889,27 +5834,66 @@ Proof MP_TAC (Q.SPECL [‘X’, ‘(M :num -> term) i’, ‘p’, ‘r’] subterm_length_first) \\ simp [Abbr ‘n’]) >> DISCH_TAC - >> qabbrev_tac ‘Y = BIGUNION (IMAGE FV (set Ms))’ - >> ‘FINITE Y’ by (rw [Abbr ‘Y’] >> REWRITE_TAC [FINITE_FV]) - (* ‘vs’ excludes all free variables in M + >> qabbrev_tac ‘d = MAX_LIST (MAP (\e. subterm_width e p) Ms)’ + >> Know ‘!i. i < k ==> subterm_width (M i) p <= d’ + >- (rw [Abbr ‘d’] \\ + MATCH_MP_TAC MAX_LIST_PROPERTY >> rw [MEM_MAP] \\ + Q.EXISTS_TAC ‘M i’ >> rw [EL_MEM, Abbr ‘M’]) + >> DISCH_TAC + >> qabbrev_tac ‘d_max = d + n_max’ + (* ‘vs0’ excludes all free variables in Ms but is in ROW 0, then is divideded + to vs and xs for constructing p1 and p3, resp. - NOTE: The basic requirement for ‘vs’ is that it must be disjoint with ‘Y’ + NOTE: The basic requirement of ‘vs’ is that it must be disjoint with ‘Y’ and is at row 0. But if we exclude ‘X UNION Y’, then it also holds that ‘set vs SUBSET X UNION RANK r’ just like another part of ‘M’. *) - >> Q_TAC (NEWS_TAC (“vs :string list”, “n_max :num”)) ‘X UNION Y’ - >> Know ‘set vs SUBSET X UNION RANK (SUC r)’ - >- (Suff ‘set vs SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ + >> Q_TAC (NEWS_TAC (“vs0 :string list”, “n_max + SUC d_max + k”)) ‘X UNION Y’ + >> Know ‘set vs0 SUBSET X UNION RANK (SUC r)’ + >- (Suff ‘set vs0 SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + qunabbrev_tac ‘vs0’ \\ MATCH_MP_TAC RNEWS_SUBSET_RANK >> simp []) >> DISCH_TAC + >> qabbrev_tac ‘vs = TAKE n_max vs0’ + >> qabbrev_tac ‘xs = DROP n_max vs0’ + >> ‘vs ++ xs = vs0’ by METIS_TAC [TAKE_DROP] + >> Know ‘set vs SUBSET set vs0 /\ set xs SUBSET set vs0’ + >- (POP_ASSUM (REWRITE_TAC o wrap o SYM) \\ + simp [LIST_TO_SET_APPEND]) + >> STRIP_TAC + >> Know ‘set vs SUBSET X UNION RANK (SUC r)’ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs0’ >> art []) + >> DISCH_TAC + >> Know ‘DISJOINT (set vs) X /\ DISJOINT (set vs) Y /\ + DISJOINT (set xs) X /\ DISJOINT (set xs) Y’ + >- (rpt CONJ_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art []) + >> STRIP_TAC + >> ‘LENGTH vs = n_max’ by rw [Abbr ‘vs’] + >> ‘LENGTH xs = SUC d_max + k’ by rw [Abbr ‘xs’] + >> Know ‘ALL_DISTINCT vs /\ ALL_DISTINCT xs /\ DISJOINT (set xs) (set vs)’ + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs0’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs ++ xs = vs0’ (REWRITE_TAC o wrap o SYM) \\ + simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) + >> STRIP_TAC + >> ‘NEWS n_max (X UNION Y) = vs’ by simp [Abbr ‘vs’, Abbr ‘vs0’, TAKE_RNEWS] + >> Know ‘set vs SUBSET ROW 0’ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs0’ >> art [] \\ + rw [Abbr ‘vs0’, RNEWS_SUBSET_ROW]) + >> DISCH_TAC + >> Know ‘set vs SUBSET RANK r’ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ >> art [] \\ + rw [ROW_SUBSET_RANK]) + >> DISCH_TAC (* construct p1 *) >> qabbrev_tac ‘p1 = MAP rightctxt (REVERSE (MAP VAR vs))’ >> ‘Boehm_transform p1’ by rw [Abbr ‘p1’, MAP_MAP_o, GSYM MAP_REVERSE] (* decompose M0 by hnf_cases_shared *) >> Know ‘!i. i < k ==> ?y args. M0 i = LAMl (TAKE (n i) vs) (VAR y @* args)’ - >- (rw [Abbr ‘n’] >> irule (iffLR hnf_cases_shared) \\ - rw [] >- fs [o_DEF] \\ + >- (Q.X_GEN_TAC ‘i’ >> DISCH_TAC \\ + fs [Abbr ‘n’] \\ + irule (iffLR hnf_cases_shared) >> simp [] \\ MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘FV (EL i Ms)’ \\ reverse CONJ_TAC @@ -5927,18 +5911,21 @@ Proof (* define M1 *) >> qabbrev_tac ‘M1 = \i. principle_hnf (M0 i @* MAP VAR vs)’ >> Know ‘!i. i < k ==> M1 i = VAR (y i) @* args i @* DROP (n i) (MAP VAR vs)’ - >- (rw [Abbr ‘M1’] \\ + >- (Q.X_GEN_TAC ‘i’ >> DISCH_TAC \\ + simp [Abbr ‘M1’] \\ qabbrev_tac ‘t = VAR (y i) @* args i’ \\ - rw [GSYM MAP_DROP] \\ - qabbrev_tac ‘xs = TAKE (n i) vs’ \\ - Know ‘MAP VAR vs = MAP VAR xs ++ MAP VAR (DROP (n i) vs)’ + simp [GSYM MAP_DROP] \\ + qabbrev_tac ‘vs' = TAKE (n i) vs’ \\ + Know ‘MAP VAR vs = MAP VAR vs' ++ MAP VAR (DROP (n i) vs)’ >- (REWRITE_TAC [GSYM MAP_APPEND] >> AP_TERM_TAC \\ - rw [Abbr ‘xs’, TAKE_DROP]) >> Rewr' \\ + rw [Abbr ‘vs'’, TAKE_DROP]) >> Rewr' \\ REWRITE_TAC [appstar_APPEND] \\ qabbrev_tac ‘l = MAP VAR (DROP (n i) vs)’ \\ MATCH_MP_TAC principle_hnf_beta_reduce_ext \\ rw [Abbr ‘t’, hnf_appstar]) >> DISCH_TAC + >> ‘!i. i < k ==> hnf_headvar (M1 i) = y i’ + by rw [hnf_head_hnf, GSYM appstar_APPEND] (* calculating ‘apply p1 (M0 i)’ *) >> Know ‘!i. i < k ==> apply p1 (M0 i) == M1 i’ >- (rw [Abbr ‘p1’, Boehm_apply_MAP_rightctxt'] \\ @@ -5949,12 +5936,6 @@ Proof rw [GSYM MAP_TAKE]) >> DISCH_TAC >> qabbrev_tac ‘m = \i. LENGTH (args i)’ - >> qabbrev_tac ‘d = MAX_LIST (MAP (\e. subterm_width e p) Ms)’ - >> Know ‘!i. i < k ==> subterm_width (M i) p <= d’ - >- (rw [Abbr ‘d’] \\ - MATCH_MP_TAC MAX_LIST_PROPERTY >> rw [MEM_MAP] \\ - Q.EXISTS_TAC ‘M i’ >> rw [EL_MEM, Abbr ‘M’]) - >> DISCH_TAC >> Know ‘!i. i < k ==> m i <= d’ >- (RW_TAC std_ss [] \\ Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M i) p’ \\ @@ -5963,75 +5944,122 @@ Proof >> DISCH_TAC (* NOTE: Thus P(d) is not enough to cover M1, whose ‘hnf_children_size’ is bounded by ‘d + n_max’. *) - >> qabbrev_tac ‘d_max = d + n_max’ - >> qabbrev_tac ‘P = permutator d_max’ + >> qabbrev_tac ‘P = \i. permutator (d_max + i)’ (* construct p2 *) - >> qabbrev_tac ‘p2 = GENLIST (\i. [P/y i]) k’ + >> qabbrev_tac ‘p2 = REVERSE (GENLIST (\i. [P i/y i]) k)’ >> ‘Boehm_transform p2’ by rw [Boehm_transform_def, Abbr ‘p2’, EVERY_GENLIST] - (* p2 can be rewritten to an ISUB *) - >> qabbrev_tac ‘sub = \k. REVERSE (GENLIST (\i. (P,y i)) k)’ + (* p2 can be rewritten to an ISUB + + NOTE: It's important to make ‘sub’ in increasing ‘P i’ for future uses. + *) + >> qabbrev_tac ‘sub = \k. GENLIST (\i. (P i,y i)) k’ >> Know ‘!t. apply p2 t = t ISUB sub k’ >- (simp [Abbr ‘p2’, Abbr ‘sub’] \\ Q.SPEC_TAC (‘k’, ‘j’) \\ Induct_on ‘j’ >- rw [] \\ - rw [GENLIST, REVERSE_SNOC]) + rw [GENLIST, REVERSE_SNOC, ISUB_SNOC]) >> DISCH_TAC (* properties of ‘sub’ (iterated substitution) *) >> Know ‘!j. DOM (sub j) = IMAGE y (count j) /\ FVS (sub j) = {}’ >- (simp [Abbr ‘sub’] \\ Induct_on ‘j’ >- rw [DOM_DEF, FVS_DEF] \\ - rw [GENLIST, REVERSE_SNOC, DOM_DEF, FVS_DEF, COUNT_SUC] >- SET_TAC [] \\ + rw [GENLIST, REVERSE_SNOC, DOM_DEF, FVS_DEF, COUNT_SUC, DOM_SNOC, FVS_SNOC] + >- SET_TAC [] \\ rw [Abbr ‘P’, FV_permutator]) >> DISCH_TAC - >> Know ‘!j t. t IN DOM (sub j) ==> VAR t ISUB (sub j) = P’ - >- (Induct_on ‘j’ >- rw [Abbr ‘sub’] \\ + (* NOTE: There may be duplicated names among all ‘y i’. The function f maps + i to the least j such that y j = y i, and P j is the ISUB result. + *) + >> Know ‘!i t. i <= k /\ t IN DOM (sub i) ==> + VAR t ISUB (sub i) = P (LEAST j. y j = t)’ + >- (Induct_on ‘i’ >- rw [Abbr ‘sub’] \\ rw [] \\ - Know ‘P ISUB sub j = P’ - >- (Q.SPEC_TAC (‘j’, ‘j'’) \\ - Induct_on ‘j'’ >> fs [GENLIST, REVERSE_SNOC, Abbr ‘sub’] \\ - Suff ‘[P/y j'] P = P’ >- rw [] \\ - MATCH_MP_TAC lemma14b >> rw [Abbr ‘P’, FV_permutator]) >> DISCH_TAC \\ - reverse (Cases_on ‘y x IN DOM (sub j)’) - >- (POP_ASSUM MP_TAC >> simp [] \\ - DISCH_THEN (MP_TAC o (Q.SPEC ‘x’)) >> rw [] \\ - ‘x = j’ by rw [] >> POP_ORW \\ - rfs [Abbr ‘sub’, GENLIST, REVERSE_SNOC]) \\ - rfs [Abbr ‘sub’, GENLIST, REVERSE_SNOC] \\ - rename1 ‘y x = y z’ \\ - Cases_on ‘y z = y j’ >> rw [] \\ - FIRST_X_ASSUM MATCH_MP_TAC \\ - Q.EXISTS_TAC ‘z’ >> art []) - >> DISCH_THEN (STRIP_ASSUME_TAC o Q.SPEC ‘k’) - >> Q.PAT_X_ASSUM ‘!j. DOM (sub j) = IMAGE y (count j) /\ FVS (sub j) = {}’ + Know ‘!i j. P i ISUB sub j = P i’ + >- (rw [Abbr ‘sub’] \\ + MATCH_MP_TAC ISUB_unchanged >> simp [Abbr ‘P’]) >> DISCH_TAC \\ + ‘sub (SUC i) = GENLIST (\i. (P i,y i)) (SUC i)’ by rw [] >> POP_ORW \\ + REWRITE_TAC [GENLIST] \\ + simp [DOM_SNOC, ISUB_SNOC, IN_UNION] \\ + Cases_on ‘y x IN DOM (sub i)’ + >- (Q.PAT_X_ASSUM ‘!t. i <= k /\ t IN DOM (sub i) ==> _’ + (MP_TAC o Q.SPEC ‘y (x :num)’) >> rw [] \\ + MATCH_MP_TAC lemma14b >> simp [Abbr ‘P’, FV_permutator]) \\ + Know ‘VAR (y x) ISUB sub i = VAR (y x)’ + >- (MATCH_MP_TAC ISUB_unchanged \\ + Q.PAT_X_ASSUM ‘!j. DOM (sub j) = _ /\ _’ K_TAC \\ + simp [DISJOINT_ALT]) >> Rewr' \\ + POP_ASSUM MP_TAC \\ + simp [] >> DISCH_TAC \\ + Know ‘x = i’ + >- (POP_ASSUM (MP_TAC o Q.SPEC ‘x’) >> rw []) \\ + DISCH_THEN (fs o wrap) >> T_TAC \\ + LEAST_ELIM_TAC \\ + CONJ_TAC >- (Q.EXISTS_TAC ‘i’ >> rw []) \\ + Q.X_GEN_TAC ‘j’ >> rw [Abbr ‘P’] \\ + CCONTR_TAC \\ + ‘i < j \/ j < i’ by rw [] \\ (* 2 subgoals, same tactics *) + METIS_TAC []) + >> DISCH_TAC + >> POP_ASSUM (STRIP_ASSUME_TAC o SIMP_RULE arith_ss [] o Q.SPEC ‘k’) + >> Q.PAT_X_ASSUM ‘!j. DOM (sub j) = IMAGE y (count j) /\ _’ (STRIP_ASSUME_TAC o Q.SPEC ‘k’) >> qabbrev_tac ‘ss = sub k’ - (* ss properties without induction *) - >> ‘!N. MEM N (MAP FST ss) ==> N = P’ - by rw [Abbr ‘ss’, Abbr ‘sub’, MAP_REVERSE, MAP_GENLIST, o_DEF, MEM_GENLIST] - (* NOTE: Now we have a list of M1's whose children size is bounded by d_max. - In the worst case, ‘P d_max @* M1 i’ will leave d_max+1 variable bindings + (* NOTE: f is the important injection from index to index *) + >> qabbrev_tac ‘f = \i. LEAST j. y j = y i’ + >> Know ‘!i. i < k ==> f i < k /\ VAR (y i) ISUB ss = P (f i)’ + >- (rw [Abbr ‘f’] \\ + LEAST_ELIM_TAC \\ + CONJ_TAC >- (Q.EXISTS_TAC ‘i’ >> rw []) \\ + Q.X_GEN_TAC ‘j’ >> rw [] \\ + SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [NOT_LESS]) \\ + ‘i < j’ by rw [] >> METIS_TAC []) + >> DISCH_TAC + >> Know ‘!j1 j2. y j1 <> y j2 ==> f j1 <> f j2’ + >- (rpt GEN_TAC >> DISCH_TAC \\ + simp [Abbr ‘f’] \\ + LEAST_ELIM_TAC \\ + CONJ_TAC >- (Q.EXISTS_TAC ‘j1’ >> rw []) \\ + Q.X_GEN_TAC ‘j3’ >> STRIP_TAC \\ + LEAST_ELIM_TAC \\ + CONJ_TAC >- (Q.EXISTS_TAC ‘j2’ >> rw []) \\ + Q.X_GEN_TAC ‘j4’ >> STRIP_TAC \\ + CCONTR_TAC >> METIS_TAC []) + >> DISCH_TAC + (* more properties of ISUB ‘ss’ *) + >> ‘!N. MEM N (MAP FST ss) ==> ?j. j < k /\ N = P j’ + by (rw [Abbr ‘ss’, Abbr ‘sub’, MAP_REVERSE, MAP_GENLIST, MEM_GENLIST]) + (* Now we have a list of M1's whose hnf_children_size is bounded by ‘d_max’. + In the worst case, ‘P @* M1 i’ will leave ‘SUC d_max’ variable bindings at most (in this case, ‘args i = 0 /\ n i = n_max’), and to finally get a "is_ready" term, we should apply a fresh list of d_max+1 variables (l). - *) - >> Q_TAC (NEWS_TAC (“xs :string list”, “SUC d_max”)) ‘X UNION Y UNION set vs’ - (* p3 is the maximal possible fresh list to be applied after the permutator *) + + After redefining P by (\i. permutator (d_max + i), in the new worst + case ‘P (f i) @* M1 i’ will leave at most ‘SUC d_max + k’ ending variables. + + NOTE: This list ‘xs’ is now part of vs0, defined as ‘DROP n_max vs0’. + + p3 is the maximal possible fresh list to be applied after the permutator + *) >> qabbrev_tac ‘p3 = MAP rightctxt (REVERSE (MAP VAR xs))’ >> ‘Boehm_transform p3’ by rw [Abbr ‘p3’, MAP_MAP_o, GSYM MAP_REVERSE] - (* pre-final stage *) - >> Q.EXISTS_TAC ‘p3 ++ p2 ++ p1’ - >> CONJ_TAC - >- (MATCH_MP_TAC Boehm_transform_APPEND >> art [] \\ - MATCH_MP_TAC Boehm_transform_APPEND >> art []) - >> ‘Boehm_transform (p3 ++ p2 ++ p1)’ - by (rpt (MATCH_MP_TAC Boehm_transform_APPEND >> art [])) + (* additional steps for explicit construction *) + >> Q.PAT_X_ASSUM ‘Boehm_transform pi'’ MP_TAC + >> Know ‘pi' = p3 ++ p2 ++ p1’ + >- (rw [Abbr ‘pi'’, Boehm_construction_def] \\ + simp [Abbr ‘p2’, LIST_EQ_REWRITE]) + >> Rewr' + (* “Boehm_construction” is now eliminated, back to old steps *) + >> qunabbrev_tac ‘pi'’ >> qabbrev_tac ‘pi = p3 ++ p2 ++ p1’ + >> DISCH_TAC (* Boehm_transform pi *) (* NOTE: requirements for ‘Z’ 1. y i IN Z /\ BIGUNION (IMAGE FV (set (args i))) SUBSET Z 2. DISJOINT (set xs) Z 3. Z SUBSET X UNION RANK (SUC r) *) - >> qabbrev_tac ‘Z = Y UNION set vs’ (* or ‘X UNION RANK r UNION (set vs)’ *) + >> qabbrev_tac ‘Z = Y UNION set vs’ + >> ‘FINITE Z’ by rw [Abbr ‘Z’] >> ‘DISJOINT (set xs) Z’ by rw [Abbr ‘Z’, DISJOINT_UNION'] (* FV properties of the head variable y (and children args) *) >> Know ‘!i. i < k ==> y i IN Z /\ @@ -6073,16 +6101,31 @@ Proof Q.EXISTS_TAC ‘FV (M1 i)’ >> art []) >> DISCH_TAC >> Know ‘Z SUBSET X UNION RANK r’ - >- (rw [Abbr ‘Z’, UNION_SUBSET] \\ + >- (simp [Abbr ‘Z’, UNION_SUBSET] \\ Suff ‘set vs SUBSET RANK r’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> rw []) + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ \\ + rw [ROW_SUBSET_RANK]) + >> DISCH_TAC + (* A better upper bound on ‘y i’ using subterm_headvar_lemma_alt *) + >> Know ‘!i. i < k ==> y i IN Y UNION set (TAKE (n i) vs)’ + >- (rpt STRIP_TAC \\ + Know ‘FV (M i) SUBSET Y’ + >- (rw [SUBSET_DEF, Abbr ‘Y’] \\ + Q.EXISTS_TAC ‘FV (M i)’ >> art [] \\ + Q.EXISTS_TAC ‘M i’ >> simp [Abbr ‘M’, EL_MEM]) >> DISCH_TAC \\ + Suff ‘y i IN FV (M i) UNION set (TAKE (n i) vs)’ + >- (POP_ASSUM MP_TAC >> SET_TAC []) \\ + ‘y i = hnf_headvar (M1 i)’ by simp [GSYM appstar_APPEND] \\ + POP_ORW \\ + MATCH_MP_TAC subterm_headvar_lemma_alt \\ + qexistsl_tac [‘X UNION Y’, ‘0’, ‘M0 (i :num)’, ‘n_max’] >> simp [] \\ + POP_ASSUM MP_TAC >> SET_TAC []) >> DISCH_TAC >> Know ‘!i h. i < k /\ h < m i ==> FV (EL h (args i)) SUBSET X UNION RANK r’ >- (rpt STRIP_TAC \\ qabbrev_tac ‘N = EL h (args i)’ \\ Q.PAT_X_ASSUM ‘!i. i < k ==> y i IN Z /\ _’ drule \\ - rw [BIGUNION_SUBSET] \\ + simp [BIGUNION_SUBSET] >> STRIP_TAC \\ Know ‘FV N SUBSET Z’ >- (POP_ASSUM MATCH_MP_TAC \\ Q.EXISTS_TAC ‘N’ >> simp [Abbr ‘N’, EL_MEM]) \\ @@ -6091,7 +6134,7 @@ Proof POP_ASSUM (REWRITE_TAC o wrap) \\ simp [UNION_SUBSET]) >> DISCH_TAC - (* A better estimate on BIGUNION (IMAGE FV (set (args i))) *) + (* a better upper bound of BIGUNION (IMAGE FV (set (args i))) *) >> Know ‘!i. i < k ==> BIGUNION (IMAGE FV (set (args i))) SUBSET FV (M i) UNION set (TAKE (n i) vs)’ >- (rpt STRIP_TAC \\ @@ -6125,15 +6168,16 @@ Proof *) >> qabbrev_tac ‘args' = \i. MAP (\t. t ISUB ss) (args i)’ >> ‘!i. LENGTH (args' i) = m i’ by rw [Abbr ‘args'’, Abbr ‘m’] - (* NOTE: In vs0, some elements are replaced by P, thus ‘set vs0 SUBSET set vs’ *) + (* NOTE: In vs, some elements are replaced by P, thus ‘set vs SUBSET set vs’ *) >> qabbrev_tac ‘args1 = MAP (\t. t ISUB ss) (MAP VAR vs)’ >> ‘LENGTH args1 = n_max’ by rw [Abbr ‘args1’] >> Know ‘BIGUNION (IMAGE FV (set args1)) SUBSET set vs’ >- (simp [Abbr ‘args1’, LIST_TO_SET_MAP, IMAGE_IMAGE] \\ rw [SUBSET_DEF, IN_BIGUNION_IMAGE] \\ POP_ASSUM MP_TAC \\ - Cases_on ‘x' IN DOM ss’ >- simp [Abbr ‘P’, FV_permutator] \\ - Know ‘VAR x' ISUB ss = VAR x'’ + rename1 ‘MEM v vs’ \\ + Cases_on ‘v IN DOM ss’ >- simp [Abbr ‘P’, FV_permutator] \\ + Know ‘VAR v ISUB ss = VAR v’ >- (MATCH_MP_TAC ISUB_VAR_FRESH' >> art []) >> Rewr' \\ simp []) >> DISCH_TAC @@ -6146,61 +6190,100 @@ Proof rw [Abbr ‘args2’, LIST_TO_SET_DROP]) >> DISCH_TAC (* calculating ‘apply p2 (M1 i)’ *) - >> ‘!i. i < k ==> apply p2 (M1 i) = P @* args' i @* args2 i’ + >> ‘!i. i < k ==> apply p2 (M1 i) = P (f i) @* args' i @* args2 i’ by rw [Abbr ‘args'’, Abbr ‘args1’, Abbr ‘args2’, GSYM appstar_APPEND, MAP_APPEND, appstar_ISUB, MAP_DROP] (* calculating ‘apply p2 ...’ until reaching a hnf *) - >> Know ‘!i. i < k ==> apply p3 (P @* args' i @* args2 i) = - P @* args' i @* args2 i @* MAP VAR xs’ + >> Know ‘!i. i < k ==> apply p3 (P (f i) @* args' i @* args2 i) = + P (f i) @* args' i @* args2 i @* MAP VAR xs’ >- rw [Abbr ‘p3’, Boehm_apply_MAP_rightctxt'] (* preparing for permutator_hreduce_more (no DISCH_TAC for above Know) *) >> qabbrev_tac ‘l = \i. args' i ++ args2 i ++ MAP VAR xs’ >> ASM_SIMP_TAC bool_ss [GSYM appstar_APPEND, APPEND_ASSOC] (* now l appears in the goal *) >> REWRITE_TAC [appstar_APPEND] - >> ‘!i. LENGTH (l i) = m i + (n_max - n i) + SUC d_max’ + (* stage work *) + >> ‘!i. LENGTH (l i) = m i + (n_max - n i) + SUC d_max + k’ by rw [Abbr ‘l’, Abbr ‘m’, Abbr ‘args2’, Abbr ‘d_max’, MAP_DROP] - >> ‘!i. d_max < LENGTH (l i)’ by rw [] - (* applying TAKE_DROP_SUC to break l into 3 pieces *) - >> MP_TAC (Q.GEN ‘i’ (ISPECL [“d_max :num”, “l (i :num) :term list”] - (GSYM TAKE_DROP_SUC))) >> simp [] - >> Rewr' + >> ‘!i. d_max + k < LENGTH (l i)’ by rw [] + >> DISCH_TAC + (* applying TAKE_DROP_SUC to break l into 3 pieces + + NOTE: New the segmentation of ‘l’ also depends on ‘i’. + *) + >> qabbrev_tac ‘d_max' = \i. d_max + f i’ + >> Know ‘!i. i < k ==> d_max' i < d_max + k’ + >- (rw [Abbr ‘d_max'’] \\ + Q_TAC (TRANS_TAC LESS_TRANS) ‘d_max + k’ >> simp []) + >> DISCH_TAC + >> Know ‘!i. i < k ==> d_max' i <= LENGTH (l i)’ + >- (rpt STRIP_TAC \\ + Q.PAT_X_ASSUM ‘!i. LENGTH (l i) = _’ K_TAC \\ + simp [Abbr ‘d_max'’] \\ + MATCH_MP_TAC LT_IMP_LE \\ + Q_TAC (TRANS_TAC LESS_TRANS) ‘d_max + k’ >> simp []) + >> DISCH_TAC + >> Know ‘!i. i < k ==> + apply p3 (P (f i) @* args' i @* args2 i) = + P (f i) @* (TAKE (d_max' i) (l i) ++ [EL (d_max' i) (l i)] ++ + DROP (SUC (d_max' i)) (l i))’ + >- (RW_TAC std_ss [] \\ + Suff ‘TAKE (d_max' i) (l i) ++ [EL (d_max' i) (l i)] ++ + DROP (SUC (d_max' i)) (l i) = l i’ >- Rewr \\ + MATCH_MP_TAC TAKE_DROP_SUC \\ + Q.PAT_X_ASSUM ‘!i. LENGTH (l i) = _ + k’ K_TAC \\ + Q_TAC (TRANS_TAC LESS_TRANS) ‘d_max + k’ >> simp []) + >> Q.PAT_X_ASSUM ‘!i. i < k ==> _ = P (f i) @* l i’ K_TAC >> REWRITE_TAC [appstar_APPEND, appstar_SING] - (* The segmentation of list l(i) - apply (p3 ++ p2 ++ p1) (M i) - - |<-- m(i)<= d -->|<-- n_max-n(i) -->|<-------------- SUC d_max -------------->| - |----- args' ----+----- args2 ------+-------------- MAP VAR xs ---------------| - |------------------------------------ l --------------------------------------| - | |<-j->| - |<-------- d_max = d + n_max ---------->| b - |------------------- Ns ----------------+-+--------------- tl ----------------| - |<-------------- d_max+1 ---------------->| + (* NOTE: The segmentation of list l(i) - apply (p3 ++ p2 ++ p1) (M i) + + |<-- m(i)<= d -->|<-- n_max-n(i) -->|<-------------- SUC d_max + k--------->| + |----- args' ----+----- args2 ------+-------------- MAP VAR xs -------------| + |------------------------------------ l ------------------------------------| + | |<-j->| + |<--- d_max + f = d + n_max + f(i) ---->|b| + |------------------- Ns(i) -------------+-+<-------------- tl(i) ---------->| + |<-------------- d_max + k + 1 ---------->| *) - >> qabbrev_tac ‘Ns = \i. TAKE d_max (l i)’ - >> qabbrev_tac ‘B = \i. EL d_max (l i)’ - >> simp [] (* this put Ns and B in use *) - >> qabbrev_tac ‘j = \i. d_max - LENGTH (args' i ++ args2 i)’ - >> ‘!i. j i < LENGTH xs’ by rw [Abbr ‘j’, Abbr ‘args'’, Abbr ‘args2’, Abbr ‘d_max’] + >> qabbrev_tac ‘Ns = \i. TAKE (d_max' i) (l i)’ + >> qabbrev_tac ‘B = \i. EL (d_max' i) (l i)’ + >> qabbrev_tac ‘tl = \i. DROP (SUC (d_max' i)) (l i)’ + >> simp [] (* this put Ns, B and tl in use *) + (* What is j here? The purpose of j is to show that (B i) is a fresh name in + xs. This j is the offset (d_max' i) of l, translated to offset of xs. *) + >> qabbrev_tac ‘j = \i. d_max' i - LENGTH (args' i ++ args2 i)’ + (* show that (j i) is valid index of xs *) + >> Know ‘!i. i < k ==> LENGTH (args' i ++ args2 i) <= d_max' i’ + >- (rpt STRIP_TAC \\ + simp [Abbr ‘j’, Abbr ‘args'’, Abbr ‘args2’, Abbr ‘d_max'’, Abbr ‘d_max’] \\ + MATCH_MP_TAC LESS_EQ_LESS_EQ_MONO >> simp []) + >> DISCH_TAC + >> Know ‘!i. i < k ==> j i < LENGTH xs’ + >- (rw [Abbr ‘j’, Abbr ‘args'’, Abbr ‘args2’, Abbr ‘d_max'’, Abbr ‘d_max’] \\ + ‘f i < k’ by rw [] \\ + simp [ADD1]) + >> DISCH_TAC >> Know ‘!i. i < k ==> ?b. EL (j i) xs = b /\ B i = VAR b’ >- (rw [Abbr ‘B’, Abbr ‘l’] \\ - Suff ‘EL d_max (args' i ++ args2 i ++ MAP VAR xs) = EL (j i) (MAP VAR xs)’ + Suff ‘EL (d_max' i) (args' i ++ args2 i ++ MAP VAR xs) = EL (j i) (MAP VAR xs)’ >- rw [EL_MAP] \\ SIMP_TAC bool_ss [Abbr ‘j’] \\ MATCH_MP_TAC EL_APPEND2 \\ - rw [Abbr ‘args'’, Abbr ‘args2’, Abbr ‘d_max’, MAP_DROP] \\ - MATCH_MP_TAC LESS_EQ_LESS_EQ_MONO >> rw [] \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> m i <= d’ MP_TAC \\ - simp [Abbr ‘m’]) + ‘f i < k’ by rw [] \\ + rw [Abbr ‘args'’, Abbr ‘args2’, Abbr ‘d_max'’, Abbr ‘d_max’, MAP_DROP] \\ + MATCH_MP_TAC LESS_EQ_LESS_EQ_MONO >> rw []) >> simp [EXT_SKOLEM_THM'] >> DISCH_THEN (Q.X_CHOOSE_THEN ‘b’ STRIP_ASSUME_TAC) (* this asserts ‘b’ *) - >> qabbrev_tac ‘tl = \i. DROP (SUC d_max) (l i)’ >> simp [] >> DISCH_TAC (* store ‘!i. i < k ==> apply p3 ...’ *) (* applying permutator_hreduce_more, it clearly reduces to a hnf *) >> Know ‘!i. i < k ==> - P @* Ns i @@ VAR (b i) @* tl i -h->* VAR (b i) @* Ns i @* tl i’ + P (f i) @* Ns i @@ VAR (b i) @* tl i -h->* VAR (b i) @* Ns i @* tl i’ >- (RW_TAC std_ss [Abbr ‘P’] \\ - MATCH_MP_TAC permutator_hreduce_more >> rw [Abbr ‘Ns’]) + MATCH_MP_TAC permutator_hreduce_more >> rw [Abbr ‘Ns’, Abbr ‘d_max'’] \\ + ‘f i < k’ by rw [] \\ + ‘d_max + f i <= LENGTH (l i)’ by rw [] \\ + simp [LENGTH_TAKE]) >> DISCH_TAC >> Know ‘!i. i < k ==> apply pi (M i) == VAR (b i) @* Ns i @* tl i’ >- (rpt STRIP_TAC \\ @@ -6239,8 +6322,7 @@ Proof now on, both ‘apply pi M == ...’ and ‘principle_hnf (apply pi M) = ...’ are simplified in parallel. *) - Q.PAT_X_ASSUM ‘!i. i < k ==> apply pi (M i) == _’ - (fn th => MP_TAC (MATCH_MP th (ASSUME “i < (k :num)”))) \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> apply pi (M i) == _’ drule \\ Q.PAT_X_ASSUM ‘Boehm_transform pi’ K_TAC \\ Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC \\ (* NOTE: No need to unabbrev ‘p2’ here since ‘apply p2 t = t ISUB sub k’ *) @@ -6252,23 +6334,22 @@ Proof DISCH_TAC (* store ‘M i @* MAP VAR vs @* MAP VAR xs ISUB sub k == ...’ *) \\ (* rewriting RHS to principle_hnf of ISUB *) Know ‘VAR (b i) @* Ns i @* tl i = - principle_hnf (P @* args' i @* args2 i @* MAP VAR xs)’ + principle_hnf (P (f i) @* args' i @* args2 i @* MAP VAR xs)’ >- (ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ ‘hnf (VAR (b i) @* Ns i @* tl i)’ by rw [GSYM appstar_APPEND, hnf_appstar] \\ - Suff ‘solvable (P @* args' i @* args2 i @* MAP VAR xs)’ + Suff ‘solvable (P (f i) @* args' i @* args2 i @* MAP VAR xs)’ >- (METIS_TAC [principle_hnf_thm']) \\ Suff ‘solvable (VAR (b i) @* Ns i @* tl i) /\ - P @* Ns i @@ VAR (b i) @* tl i == VAR (b i) @* Ns i @* tl i’ + P (f i) @* Ns i @@ VAR (b i) @* tl i == VAR (b i) @* Ns i @* tl i’ >- (METIS_TAC [lameq_solvable_cong]) \\ reverse CONJ_TAC >- (MATCH_MP_TAC hreduces_lameq >> rw []) \\ REWRITE_TAC [solvable_iff_has_hnf] \\ MATCH_MP_TAC hnf_has_hnf >> art []) >> Rewr' \\ - Know ‘P @* args' i @* args2 i @* MAP VAR xs = M1 i @* MAP VAR xs ISUB ss’ + Know ‘P (f i) @* args' i @* args2 i @* MAP VAR xs = M1 i @* MAP VAR xs ISUB ss’ >- (REWRITE_TAC [appstar_ISUB, Once EQ_SYM_EQ] \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> apply p2 (M1 i) = P @* args' i @* args2 i’ - (fn th => MP_TAC (MATCH_MP (GSYM (Q.SPEC ‘i’ th)) - (ASSUME “i < k :num”))) >> Rewr' \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> apply p2 (M1 i) = _’ + (drule o GSYM) >> Rewr' \\ Q.PAT_X_ASSUM ‘!t. apply p2 t = t ISUB ss’ (ONCE_REWRITE_TAC o wrap) \\ AP_TERM_TAC >> art []) >> Rewr' \\ (* applying principle_hnf_ISUB_cong *) @@ -6336,28 +6417,28 @@ Proof REWRITE_TAC [solvable_iff_has_hnf] \\ MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar, GSYM appstar_APPEND]) >> DISCH_TAC - >> CONJ_TAC (* EVERY is_ready' ... *) + >> PRINT_TAC "stage work on subtree_equiv_lemma: L6433" + >> CONJ_TAC (* EVERY is_ready ... *) >- (rpt (Q.PAT_X_ASSUM ‘Boehm_transform _’ K_TAC) \\ simp [EVERY_EL, EL_MAP] \\ Q.X_GEN_TAC ‘i’ >> DISCH_TAC \\ (* now expanding ‘is_ready’ using [is_ready_alt] *) - ASM_SIMP_TAC std_ss [is_ready_alt'] \\ + ASM_SIMP_TAC std_ss [is_ready_alt] \\ qexistsl_tac [‘b i’, ‘Ns i ++ tl i’] \\ (* subgoal: apply pi (M i) -h->* VAR (b i) @* (Ns i ++ tl i) *) CONJ_TAC >- (REWRITE_TAC [appstar_APPEND] \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> principle_hnf (apply pi (M i)) = _’ - (fn th => MP_TAC (MATCH_MP th (ASSUME “i < k:num”))) \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> principle_hnf (apply pi (M i)) = _’ drule \\ rw [principle_hnf_thm']) \\ (* final goal (is_ready): EVERY (\e. b # e) ... *) Q.PAT_X_ASSUM ‘!i. i < k ==> principle_hnf (apply pi (M i)) = _’ K_TAC \\ ASM_SIMP_TAC list_ss [EVERY_EL] \\ (* easier goal first *) reverse CONJ_TAC (* b i # EL n (tl i) *) - >- (Q.X_GEN_TAC ‘a’ >> STRIP_TAC \\ + >- (Q.X_GEN_TAC ‘a’ >> DISCH_TAC \\ qabbrev_tac ‘l1 = args' i ++ args2 i’ \\ - Know ‘LENGTH l1 <= d_max’ - >- (rw [Abbr ‘l1’, Abbr ‘args2’, Abbr ‘d_max’, MAP_DROP] \\ + Know ‘LENGTH l1 <= d_max' i’ + >- (rw [Abbr ‘l1’, Abbr ‘args2’, Abbr ‘d_max’, Abbr ‘d_max'’, MAP_DROP] \\ MATCH_MP_TAC LESS_EQ_LESS_EQ_MONO >> rw [] \\ Q.PAT_X_ASSUM ‘!i. i < k ==> m i <= d’ MP_TAC \\ simp [Abbr ‘m’]) >> DISCH_TAC \\ @@ -6368,9 +6449,9 @@ Proof (* applying DROP_APPEND2 *) Know ‘tl i = DROP (SUC (j i)) (MAP VAR xs)’ >- (rw [Abbr ‘tl’, Abbr ‘l’] \\ - ‘LENGTH l1 <= SUC d_max’ by rw [] \\ + ‘LENGTH l1 <= SUC (d_max' i)’ by rw [] \\ ASM_SIMP_TAC std_ss [DROP_APPEND2] \\ - Suff ‘SUC d_max - LENGTH l1 = SUC (j i)’ >- rw [] \\ + Suff ‘SUC (d_max' i) - LENGTH l1 = SUC (j i)’ >- rw [] \\ ASM_SIMP_TAC arith_ss [Abbr ‘j’]) >> Rewr' \\ simp [] >> DISCH_TAC (* a < d_max + ... *) \\ Know ‘EL a (DROP (SUC (j i)) (MAP VAR xs :term list)) = @@ -6388,15 +6469,30 @@ Proof Q.PAT_X_ASSUM ‘!i. i < k ==> apply pi (M i) == _’ K_TAC \\ Q.PAT_X_ASSUM ‘!i. i < k ==> apply p3 _ = _’ K_TAC \\ Q.PAT_X_ASSUM ‘!i. i < k ==> apply p2 _ = _’ K_TAC \\ - Q.PAT_X_ASSUM ‘!t. apply p2 t = t ISUB ss’ K_TAC \\ + Q.PAT_X_ASSUM ‘!t. apply p2 t = t ISUB ss’ K_TAC \\ Q.PAT_X_ASSUM ‘!i. i < k ==> apply p1 _ == _’ K_TAC \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> P @* Ns i @@ VAR (b i) @* tl i -h->* _’ K_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> solvable (apply pi (M i))’ K_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> P (f i) @* Ns i @@ _ @* tl i -h->* _’ K_TAC \\ qunabbrevl_tac [‘pi’, ‘p1’, ‘p2’, ‘p3’] \\ - (* first case *) + (* The segmentation of list l(i) - apply (p3 ++ p2 ++ p1) (M i) + + |<-- m(i)<= d -->|<-- n_max-n(i) -->|<------------- SUC d_max + k --------->| + |----- args' ----+----- args2 ------+-------------- MAP VAR xs -------------| + |------------------------------------ l ------------------------------------| + | |<-j->| + |<--- d_max + f = d + n_max + f(i) ---->|b| + |------------------- Ns(i) -------------+-+<-------------- tl(i) ---------->| + |<-------------- d_max + k + 1 ---------->| + + Three cases for proving ‘b i # EL a (Ns i)’, given a < LENGTH (Ns i): + 1) a < m i (= LENGTH (args' i)) + 2) m i <= a < m i + LENGTH (args2 i) + 3) m i + LENGTH (args2 i) <= a + *) Cases_on ‘a < m i’ >- (Q.PAT_X_ASSUM ‘a < LENGTH (Ns i)’ MP_TAC \\ simp [Abbr ‘Ns’, LENGTH_TAKE] \\ - DISCH_TAC (* a < d_max *) \\ + DISCH_TAC (* a < d_max' i *) \\ simp [EL_TAKE] \\ Know ‘EL a (l i) = EL a (args' i)’ >- (SIMP_TAC std_ss [Abbr ‘l’, GSYM APPEND_ASSOC] \\ @@ -6480,8 +6576,8 @@ Proof SPOSE_NOT_THEN (STRIP_ASSUME_TAC o REWRITE_RULE []) \\ Suff ‘EL (j i) xs = EL a' xs <=> j i = a'’ >- rw [] \\ MATCH_MP_TAC ALL_DISTINCT_EL_IMP >> rw []) - >> PRINT_TAC "[subterm_equiv_lemma] stage work" (* cleanup *) + >> PRINT_TAC "stage work on subtree_equiv_lemma: L6593" >> Q.PAT_X_ASSUM ‘Boehm_transform p1’ K_TAC >> Q.PAT_X_ASSUM ‘Boehm_transform p2’ K_TAC >> Q.PAT_X_ASSUM ‘Boehm_transform p3’ K_TAC @@ -6513,10 +6609,10 @@ Proof Know ‘DISJOINT (set vs) (set ys')’ >- (qunabbrev_tac ‘ys'’ \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘RANK r’ >> rw [DISJOINT_RANK_RNEWS'] \\ - Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ \\ - CONJ_TAC >- rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ - MATCH_MP_TAC ROW_SUBSET_RANK >> art []) >> DISCH_TAC \\ + Q.EXISTS_TAC ‘RANK r’ \\ + simp [DISJOINT_RANK_RNEWS'] \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ >> art [] \\ + rw [ROW_SUBSET_RANK]) >> DISCH_TAC \\ Know ‘DISJOINT (set vs') (set ys')’ >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ >> art [] \\ @@ -6536,9 +6632,14 @@ Proof simp [ltree_lookup, LMAP_fromList, MAP_MAP_o, LNTH_fromList, EL_MAP] \\ Cases_on ‘h < m i’ >> simp [] \\ qabbrev_tac ‘pm = ZIP (vs',ys')’ \\ + Know ‘d_max' i <= LENGTH (l i)’ + >- (Q.PAT_X_ASSUM ‘!i. LENGTH (l i) = _’ K_TAC \\ + simp [Abbr ‘d_max'’] \\ + MATCH_MP_TAC LT_IMP_LE \\ + Q_TAC (TRANS_TAC LESS_TRANS) ‘d_max + k’ >> simp []) >> DISCH_TAC \\ Know ‘h < LENGTH (Ns i)’ >- (simp [Abbr ‘Ns’] \\ - Suff ‘m i <= d_max’ >- rw [] \\ + Suff ‘m i <= d_max’ >- rw [Abbr ‘d_max'’] \\ simp [Abbr ‘d_max’] \\ MATCH_MP_TAC LESS_EQ_TRANS \\ Q.EXISTS_TAC ‘d’ >> simp []) >> DISCH_TAC \\ @@ -6565,8 +6666,8 @@ Proof impl_tac >- (Q.EXISTS_TAC ‘EL h (args i)’ >> rw [EL_MEM]) \\ Q.PAT_X_ASSUM ‘lswapstr (REVERSE pm) x IN FV (EL h (args i))’ K_TAC \\ Know ‘set vs SUBSET RANK (SUC r)’ - >- (qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> rw []) >> DISCH_TAC \\ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ >> art [] \\ + rw [ROW_SUBSET_RANK]) >> DISCH_TAC \\ Know ‘set vs' SUBSET RANK (SUC r)’ >- (qunabbrev_tac ‘vs'’ \\ Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs’ \\ @@ -6576,7 +6677,7 @@ Proof MATCH_MP_TAC RNEWS_SUBSET_RANK >> rw []) >> DISCH_TAC \\ simp [Abbr ‘Z’, IN_UNION] \\ reverse STRIP_TAC - (* lswapstr (REVERSE pm) x IN set vs *) + (* when “lswapstr (REVERSE pm) x IN set vs” is assumed *) >- (DISJ2_TAC >> POP_ASSUM MP_TAC >> simp [MEM_EL] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘a’ STRIP_ASSUME_TAC) \\ Know ‘x = lswapstr pm (EL a vs)’ @@ -6587,14 +6688,11 @@ Proof rw [SUBSET_DEF, MEM_EL] \\ POP_ASSUM MATCH_MP_TAC \\ Q.EXISTS_TAC ‘a’ >> art []) \\ - Know ‘set vs SUBSET ROW 0’ - >- (qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_ROW >> rw []) >> DISCH_TAC \\ Know ‘set ys' SUBSET ROW r’ >- (qunabbrev_tac ‘ys'’ \\ MATCH_MP_TAC RNEWS_SUBSET_ROW >> art []) >> DISCH_TAC \\ Know ‘DISJOINT (ROW 0) (ROW r)’ >- rw [ROW_DISJOINT] \\ - rw [DISJOINT_ALT] \\ + simp [DISJOINT_ALT] >> STRIP_TAC \\ Suff ‘EL a vs NOTIN ROW r’ >- METIS_TAC [SUBSET_DEF] \\ FIRST_X_ASSUM MATCH_MP_TAC \\ Suff ‘EL a vs IN set vs’ >- METIS_TAC [SUBSET_DEF] \\ @@ -6602,14 +6700,16 @@ Proof (* lswapstr (REVERSE pm) x IN Y (SUBSET X UNION RANK r) *) Know ‘lswapstr (REVERSE pm) x IN X UNION RANK r’ >- ASM_SET_TAC [] \\ Q.PAT_X_ASSUM ‘lswapstr (REVERSE pm) x IN Y’ K_TAC \\ - RW_TAC std_ss [IN_UNION] + simp [IN_UNION] \\ + STRIP_TAC >- (FULL_SIMP_TAC std_ss [GSYM ssetpm_IN] \\ DISJ1_TAC \\ Suff ‘ssetpm pm X = X’ >- DISCH_THEN (FULL_SIMP_TAC std_ss o wrap) \\ - MATCH_MP_TAC ssetpm_unchanged >> rw [Abbr ‘pm’, MAP_ZIP] \\ + MATCH_MP_TAC ssetpm_unchanged \\ + simp [Abbr ‘pm’, MAP_ZIP] \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ >> art [] \\ - rw [Abbr ‘vs'’, LIST_TO_SET_TAKE]) \\ + simp [Abbr ‘vs'’, LIST_TO_SET_TAKE]) \\ DISJ2_TAC \\ FULL_SIMP_TAC std_ss [GSYM ssetpm_IN] \\ qabbrev_tac ‘x' = lswapstr (REVERSE pm) x’ \\ @@ -6649,33 +6749,29 @@ Proof MATCH_MP_TAC RNEWS_SUBSET_RANK >> rw []) \\ Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs’ \\ CONJ_TAC >- rw [Abbr ‘vs'’, LIST_TO_SET_TAKE] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> rw []) \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ >> art [] \\ + rw [ROW_SUBSET_RANK]) \\ NTAC 2 (POP_ASSUM K_TAC) \\ fs [Abbr ‘pm'’, Abbr ‘N’] >> T_TAC \\ qabbrev_tac ‘N = EL h (args i)’ \\ - (* applying BT_ltree_paths_isub_cong *) + (* applying subterm_width_isub_permutator_cong *) ‘!M. ltree_lookup (BT' X M (SUC r)) t <> NONE <=> t IN ltree_paths (BT' X M (SUC r))’ by rw [ltree_paths_def] \\ POP_ORW >> DISCH_TAC \\ - MATCH_MP_TAC (cj 1 BT_ltree_paths_isub_cong) \\ - qexistsl_tac [‘REVERSE (GENLIST y k)’, ‘P’, ‘d_max’] >> simp [] \\ + MATCH_MP_TAC (cj 1 subterm_width_isub_permutator_cong_alt) \\ + qexistsl_tac [‘d_max’, ‘y’, ‘k’] >> simp [] \\ CONJ_TAC >- (Suff ‘FV N SUBSET X UNION RANK r’ >- (Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ rw [RANK_MONO]) \\ qunabbrev_tac ‘N’ >> FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ CONJ_TAC - >- (Q.X_GEN_TAC ‘v’ >> simp [MEM_GENLIST] \\ - DISCH_THEN (Q.X_CHOOSE_THEN ‘J’ STRIP_ASSUME_TAC) >> POP_ORW \\ + >- (Q.X_GEN_TAC ‘J’ >> DISCH_TAC \\ Know ‘y J IN Z’ >- rw [] \\ Suff ‘Z SUBSET X UNION RANK (SUC r)’ >- rw [SUBSET_DEF] \\ Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ >> art [] \\ Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ rw [RANK_MONO]) \\ - CONJ_TAC (* ss = MAP ... *) - >- (simp [Abbr ‘ss’, Abbr ‘sub’] \\ - simp [MAP_REVERSE, MAP_GENLIST, o_DEF]) \\ Know ‘subterm_width (M i) (h::t) <= d’ >- rw [] \\ qabbrev_tac ‘Ms' = args i ++ DROP (n i) (MAP VAR vs)’ \\ (* applying subterm_width_induction_lemma (the general one) *) @@ -6697,12 +6793,10 @@ Proof Q.EXISTS_TAC ‘M i’ >> rw [Abbr ‘M’, EL_MEM]) >> DISCH_TAC (* stage work *) - >> CONJ_TAC - >- (RW_TAC std_ss [MEM_EL] \\ - FIRST_X_ASSUM MATCH_MP_TAC >> art []) - (* An upper bound of all FVs from l (args' ++ args2 ++ xs) *) + >> CONJ_TAC >- (rw [EVERY_EL] >> simp [EL_MAP]) + (* upper bound of all FVs from l (args' ++ args2 ++ xs) *) >> Know ‘!i. i < k ==> BIGUNION (IMAGE FV (set (l i))) SUBSET X UNION RANK r’ - >- (rw [Abbr ‘l’] >| (* 3 subgoals *) + >- (simp [Abbr ‘l’] >> rpt STRIP_TAC >| (* 3 subgoals *) [ (* goal 1 (of 3): args' *) Q_TAC (TRANS_TAC SUBSET_TRANS) ‘Z’ >> art [] \\ Q_TAC (TRANS_TAC SUBSET_TRANS) ‘BIGUNION (IMAGE FV (set (args i)))’ \\ @@ -6717,7 +6811,7 @@ Proof Suff ‘FV (EL h (args' i)) SUBSET FV (EL h (args i))’ >- rw [SUBSET_DEF] \\ simp [Abbr ‘args'’, EL_MAP] \\ qabbrev_tac ‘N = EL h (args i)’ \\ - MP_TAC (Q.SPECL [‘ss’, ‘N’] FV_ISUB_upperbound) >> simp [], + MP_TAC (Q.SPECL [‘ss’, ‘N’] FV_ISUB_upperbound) >> rw [], (* goal 2 (of 3): args2 *) Q_TAC (TRANS_TAC SUBSET_TRANS) ‘BIGUNION (IMAGE FV (set args1))’ \\ CONJ_TAC @@ -6733,11 +6827,9 @@ Proof Q_TAC (TRANS_TAC SUBSET_TRANS) ‘RANK r’ \\ reverse CONJ_TAC >- SET_TAC [] \\ Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ \\ - reverse CONJ_TAC >- (MATCH_MP_TAC ROW_SUBSET_RANK >> art []) \\ - rw [IN_BIGUNION_IMAGE, SUBSET_DEF, MEM_MAP] \\ - POP_ASSUM MP_TAC >> rw [] \\ - Suff ‘set xs SUBSET ROW 0’ >- rw [SUBSET_DEF] \\ - rw [Abbr ‘xs’, RNEWS_SUBSET_ROW] ]) + simp [ROW_SUBSET_RANK] \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs0’ >> art [] \\ + simp [Abbr ‘vs0’, RNEWS_SUBSET_ROW] ]) >> DISCH_TAC (* ‘H i’ is the head reduction of apply pi (M i) *) >> qabbrev_tac ‘H = \i. VAR (b i) @* Ns i @* tl i’ @@ -6746,7 +6838,8 @@ Proof MATCH_MP_TAC hnf_has_hnf >> rw [hnf_appstar]) >> DISCH_TAC >> Know ‘!i. i < k ==> FV (H i) SUBSET X UNION RANK r’ - >- (rw [Abbr ‘H’, GSYM appstar_APPEND, FV_appstar] >| (* 3 subgoals *) + >- (simp [Abbr ‘H’, GSYM appstar_APPEND, FV_appstar] \\ + rpt STRIP_TAC >| (* 3 subgoals *) [ (* goal 1 (of 3): easier *) REWRITE_TAC [IN_UNION] >> DISJ2_TAC \\ Suff ‘b i IN ROW 0’ @@ -6758,32 +6851,34 @@ Proof Suff ‘set xs SUBSET ROW 0’ >- (rw [SUBSET_DEF] \\ POP_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) \\ - qunabbrev_tac ‘xs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_ROW >> simp [], + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘set vs0’ >> art [] \\ + simp [Abbr ‘vs0’, RNEWS_SUBSET_ROW], (* goal 2 (of 3): hard but now easy *) Q_TAC (TRANS_TAC SUBSET_TRANS) ‘BIGUNION (IMAGE FV (set (l i)))’ \\ simp [Abbr ‘Ns’] \\ MATCH_MP_TAC SUBSET_BIGUNION \\ - MATCH_MP_TAC IMAGE_SUBSET \\ - rw [LIST_TO_SET_TAKE], + MATCH_MP_TAC IMAGE_SUBSET >> rw [LIST_TO_SET_TAKE], (* goal 3 (of 3): not that hard *) Q_TAC (TRANS_TAC SUBSET_TRANS) ‘BIGUNION (IMAGE FV (set (l i)))’ \\ simp [Abbr ‘tl’] \\ MATCH_MP_TAC SUBSET_BIGUNION \\ - MATCH_MP_TAC IMAGE_SUBSET \\ - rw [LIST_TO_SET_DROP] ]) + MATCH_MP_TAC IMAGE_SUBSET >> rw [LIST_TO_SET_DROP] ]) >> DISCH_TAC >> Know ‘!i. i < k ==> d_max <= LENGTH (hnf_children (H i))’ >- (rw [Abbr ‘H’, GSYM appstar_APPEND] \\ - simp [Abbr ‘Ns’]) + simp [Abbr ‘Ns’] \\ + simp [Abbr ‘d_max'’]) >> DISCH_TAC (* stage work, now connect ‘subterm X (M i) q’ with ‘subterm X (H i) q’ *) >> Q_TAC (RNEWS_TAC (“ys :string list”, “r :num”, “n_max :num”)) ‘X’ >> qabbrev_tac ‘pm = ZIP (vs,ys)’ >> Know ‘DISJOINT (set vs) (set ys)’ - >- (qunabbrevl_tac [‘vs’, ‘ys’] \\ + >- (Q.PAT_X_ASSUM ‘_ = vs’ (REWRITE_TAC o wrap o SYM) \\ + qunabbrev_tac ‘ys’ \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp []) >> DISCH_TAC + >> PRINT_TAC "stage work on subtree_equiv_lemma: L6893" + (* Now ‘subterm X (M i) q r <> NONE’ is added into antecedents of the subgoal *) >> Know ‘!q. q <<= p /\ q <> [] ==> !i. i < k ==> subterm X (H i) q r <> NONE /\ subterm' X (H i) q r = @@ -6803,8 +6898,8 @@ Proof ‘LAMl_size (H i) = 0’ by rw [Abbr ‘H’, LAMl_size_appstar, GSYM appstar_APPEND] \\ simp [] \\ - NTAC 2 (POP_ASSUM K_TAC) \\ - DISCH_TAC \\ + NTAC 2 (POP_ASSUM K_TAC) (* principle_hnf (H i), LAMl_size (H i) *) \\ + DISCH_TAC (* subterm_width (M i) (h::t) <= d *) \\ Q_TAC (RNEWS_TAC (“ys' :string list”, “r :num”, “(n :num -> num) i”)) ‘X’ \\ qabbrev_tac ‘vs' = TAKE (n i) vs’ \\ ‘ALL_DISTINCT vs' /\ LENGTH vs' = n i’ @@ -6820,12 +6915,12 @@ Proof simp [LENGTH_TAKE]) \\ DISCH_THEN (rfs o wrap) >> T_TAC \\ Know ‘DISJOINT (set vs) (set ys)’ - >- (qunabbrev_tac ‘ys’ \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘RANK r’ >> rw [DISJOINT_RANK_RNEWS'] \\ - Q_TAC (TRANS_TAC SUBSET_TRANS) ‘ROW 0’ \\ - CONJ_TAC >- rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ - MATCH_MP_TAC ROW_SUBSET_RANK >> art []) >> DISCH_TAC \\ + >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘ROW 0’ >> art [] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘ROW r’ \\ + simp [Abbr ‘ys’, RNEWS_SUBSET_ROW] \\ + rw [Once DISJOINT_SYM, ROW_DISJOINT]) >> DISCH_TAC \\ Know ‘DISJOINT (set vs') (set ys')’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set ys’ \\ @@ -6847,6 +6942,7 @@ Proof >- (Q_TAC (TRANS_TAC LESS_LESS_EQ_TRANS) ‘m i’ >> art [] \\ Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘d’ >> simp [] \\ simp [Abbr ‘d_max’]) >> DISCH_TAC \\ + ‘h < d_max' i’ by rw [Abbr ‘d_max'’] \\ Know ‘h < LENGTH (hnf_children (H i))’ >- (Q_TAC (TRANS_TAC LESS_LESS_EQ_TRANS) ‘d_max’ \\ simp []) >> Rewr \\ @@ -6855,18 +6951,18 @@ Proof MATCH_MP_TAC EL_APPEND1 >> simp [Abbr ‘Ns’]) >> Rewr' \\ Know ‘EL h (Ns i) = EL h (args' i)’ >- (simp [Abbr ‘Ns’] \\ - Know ‘EL h (TAKE d_max (l i)) = EL h (l i)’ + Know ‘EL h (TAKE (d_max' i) (l i)) = EL h (l i)’ >- (MATCH_MP_TAC EL_TAKE >> art []) >> Rewr' \\ simp [Abbr ‘l’] \\ REWRITE_TAC [GSYM APPEND_ASSOC] \\ MATCH_MP_TAC EL_APPEND1 \\ simp [Abbr ‘args'’]) >> Rewr' \\ qabbrev_tac ‘N = EL h (args i)’ \\ - fs [Abbr ‘args'’, EL_MAP] \\ + fs [Abbr ‘args'’, EL_MAP] >> T_TAC \\ qabbrev_tac ‘pm' = ZIP (vs',ys')’ \\ Q.PAT_X_ASSUM ‘DISJOINT (set vs') (set ys')’ K_TAC \\ (* applying tpm_unchanged *) - Know ‘tpm pm' N = tpm pm N’ (* (n i) vs n_max *) + Know ‘tpm pm' N = tpm pm N’ (* (n i) vs. n_max *) >- (simp [Abbr ‘pm'’, Abbr ‘vs'’] \\ Q.PAT_X_ASSUM ‘TAKE (n i) ys = ys'’ (REWRITE_TAC o wrap o SYM) \\ simp [ZIP_TAKE] \\ @@ -6895,25 +6991,27 @@ Proof MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set ys’ \\ simp [LIST_TO_SET_DROP] \\ - rw [DISJOINT_ALT'] \\ + simp [DISJOINT_ALT'] >> rpt STRIP_TAC \\ Know ‘x IN FV (M i) UNION set (TAKE (n i) vs)’ >- METIS_TAC [SUBSET_DEF] \\ - rw [IN_UNION] + simp [IN_UNION] \\ + CONJ_TAC >- (Q.PAT_X_ASSUM ‘DISJOINT (set ys) (FV (M i))’ MP_TAC \\ - rw [DISJOINT_ALT']) \\ + rw [DISJOINT_ALT]) \\ Suff ‘DISJOINT (set (TAKE (n i) vs)) (set ys)’ - >- rw [DISJOINT_ALT] \\ + >- rw [DISJOINT_ALT'] \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ >> rw [LIST_TO_SET_TAKE]) \\ (* vs is slightly harder *) - rw [DISJOINT_ALT'] \\ + simp [DISJOINT_ALT'] >> rpt STRIP_TAC \\ Know ‘x IN FV (M i) UNION set (TAKE (n i) vs)’ >- METIS_TAC [SUBSET_DEF] \\ - reverse (rw [IN_UNION]) + simp [IN_UNION] \\ + reverse CONJ_TAC >- (Know ‘ALL_DISTINCT (TAKE (n i) vs ++ DROP (n i) vs)’ >- rw [TAKE_DROP] \\ - rw [ALL_DISTINCT_APPEND]) \\ - Suff ‘DISJOINT (set (DROP (n i) vs)) (FV (M i))’ >- rw [DISJOINT_ALT'] \\ + simp [ALL_DISTINCT_APPEND', DISJOINT_ALT']) \\ + Suff ‘DISJOINT (set (DROP (n i) vs)) (FV (M i))’ >- rw [DISJOINT_ALT] \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ \\ simp [LIST_TO_SET_DROP] \\ @@ -6932,7 +7030,7 @@ Proof Know ‘FV N SUBSET X UNION RANK (SUC r)’ >- (Suff ‘FV N SUBSET X UNION RANK r’ >- (Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) \\ + rw [RANK_MONO]) \\ qunabbrev_tac ‘N’ \\ FIRST_X_ASSUM MATCH_MP_TAC >> art []) >> DISCH_TAC \\ Know ‘DISJOINT (set ys) (FV (M i))’ @@ -6940,7 +7038,7 @@ Proof qexistsl_tac [‘X’, ‘r’, ‘n_max’] >> simp []) >> DISCH_TAC \\ Know ‘DISJOINT (set ys) (FV N)’ >- (Q.PAT_X_ASSUM ‘!i. i < k ==> _ SUBSET FV (M i) UNION _’ drule \\ - rw [BIGUNION_SUBSET] \\ + simp [BIGUNION_SUBSET] >> STRIP_TAC \\ MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘FV (M i) UNION set vs'’ \\ reverse CONJ_TAC @@ -6955,25 +7053,22 @@ Proof DISCH_TAC (* subterm X (tpm pm N) t (SUC r) <> NONE *) \\ MP_TAC (Q.SPECL [‘pm’, ‘X’, ‘N’, ‘t'’, ‘SUC r’] subterm_fresh_tpm_cong) \\ impl_tac >- simp [Abbr ‘pm’, MAP_ZIP] \\ - simp [] >> STRIP_TAC \\ - POP_ASSUM K_TAC (* already used *) \\ + simp [] \\ + STRIP_TAC >> POP_ASSUM K_TAC (* already used *) \\ (* applying subterm_isub_permutator_cong' *) - MATCH_MP_TAC subterm_isub_permutator_cong' \\ - qexistsl_tac [‘REVERSE (GENLIST y k)’, ‘P’, ‘d_max’] >> simp [] \\ - reverse CONJ_TAC (* easier *) - >- (CONJ_TAC - >- (rw [MEM_GENLIST] \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> y i IN Z /\ _’ drule \\ - qunabbrev_tac ‘Z’ >> STRIP_TAC \\ - rename1 ‘i' < k’ \\ - Q.PAT_X_ASSUM ‘y i' IN Y UNION set vs’ MP_TAC \\ - Suff ‘Y UNION set vs SUBSET X UNION RANK (SUC r)’ >- SET_TAC [] \\ - rw [UNION_SUBSET] \\ - Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ >> art [] \\ - Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ - rw [RANK_MONO]) \\ - simp [Abbr ‘ss’, Abbr ‘sub’] \\ - simp [MAP_REVERSE, MAP_GENLIST, o_DEF]) \\ + MATCH_MP_TAC subterm_isub_permutator_cong_alt' \\ + qexistsl_tac [‘d_max’, ‘y’, ‘k’] >> simp [] \\ + CONJ_TAC (* easier *) + >- (rpt STRIP_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> y i IN Z /\ _’ drule \\ + qunabbrev_tac ‘Z’ >> STRIP_TAC \\ + rename1 ‘i' < k’ \\ + Q.PAT_X_ASSUM ‘y i' IN Y UNION set vs’ MP_TAC \\ + Suff ‘Y UNION set vs SUBSET X UNION RANK (SUC r)’ >- SET_TAC [] \\ + rw [UNION_SUBSET] \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ >> art [] \\ + Suff ‘RANK r SUBSET RANK (SUC r)’ >- SET_TAC [] \\ + rw [RANK_MONO]) \\ (* subterm_width N t <= d_max *) Know ‘subterm_width (M i) (h::t') <= d’ >- (MATCH_MP_TAC subterm_width_inclusive \\ @@ -7000,6 +7095,16 @@ Proof MATCH_MP_TAC ltree_paths_inclusive \\ Q.EXISTS_TAC ‘h::t’ >> simp []) >> DISCH_TAC + >> PRINT_TAC "stage work on subtree_equiv_lemma: L7111" + >> Suff ‘(!M N q. + MEM M Ms /\ MEM N Ms /\ q <<= p /\ + subtree_equiv X M N q r ==> + subtree_equiv X (apply pi M) (apply pi N) q r) /\ + (!M N q. + MEM M Ms /\ MEM N Ms /\ q <<= p /\ + subtree_equiv X (apply pi M) (apply pi N) q r ==> + subtree_equiv X M N q r)’ + >- METIS_TAC [] (* stage work, next goal: !M N q. MEM M Ms /\ MEM N Ms /\ q <<= p /\ subtree_equiv X M N q r ==> @@ -7038,58 +7143,49 @@ Proof where the relation ~1~ is to be established by BT_subterm_thm, and ~2~ follows a similar idea of [Boehm_transform_exists_lemma]. - - The case ‘q = []’ is special: *) Cases_on ‘q = []’ >- (POP_ORW >> simp [BT_ltree_el_NIL] \\ Know ‘!i. principle_hnf (H i) = H i’ >- (rw [Abbr ‘H’] >> MATCH_MP_TAC principle_hnf_reduce \\ rw [hnf_appstar]) >> Rewr' \\ - ‘!i. LAMl_size (H i) = 0’ - by rw [Abbr ‘H’, GSYM appstar_APPEND, LAMl_size_appstar] \\ + Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ K_TAC \\ simp [Abbr ‘H’, GSYM appstar_APPEND, hnf_head_appstar] \\ simp [head_equivalent_def] \\ - qabbrev_tac ‘ys1 = TAKE (n j1) vs’ \\ - qabbrev_tac ‘ys2 = TAKE (n j2) vs’ \\ - ‘ALL_DISTINCT ys1 /\ ALL_DISTINCT ys2’ - by simp [Abbr ‘ys1’, Abbr ‘ys2’, ALL_DISTINCT_TAKE] \\ - ‘LENGTH ys1 = n j1’ - by (qunabbrev_tac ‘ys1’ \\ + qabbrev_tac ‘vs1 = TAKE (n j1) vs’ \\ + qabbrev_tac ‘vs2 = TAKE (n j2) vs’ \\ + ‘ALL_DISTINCT vs1 /\ ALL_DISTINCT vs2’ + by simp [Abbr ‘vs1’, Abbr ‘vs2’, ALL_DISTINCT_TAKE] \\ + ‘LENGTH vs1 = n j1’ + by (qunabbrev_tac ‘vs1’ \\ MATCH_MP_TAC LENGTH_TAKE >> art [] \\ FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ - ‘LENGTH ys2 = n j2’ - by (qunabbrev_tac ‘ys2’ \\ + ‘LENGTH vs2 = n j2’ + by (qunabbrev_tac ‘vs2’ \\ MATCH_MP_TAC LENGTH_TAKE >> art [] \\ FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ - Q_TAC (RNEWS_TAC (“zs1 :string list”, “r :num”, “(n :num -> num) j1”)) ‘X’ \\ - Q_TAC (RNEWS_TAC (“zs2 :string list”, “r :num”, “(n :num -> num) j2”)) ‘X’ \\ - Know ‘DISJOINT (set ys1) (set zs1)’ + Q_TAC (RNEWS_TAC (“ys1 :string list”, “r :num”, + “(n :num -> num) j1”)) ‘X’ \\ + Q_TAC (RNEWS_TAC (“ys2 :string list”, “r :num”, + “(n :num -> num) j2”)) ‘X’ \\ + Know ‘DISJOINT (set vs1) (set ys1)’ >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ \\ - reverse CONJ_TAC >- rw [Abbr ‘ys1’, LIST_TO_SET_TAKE] \\ - qunabbrev_tac ‘zs1’ \\ + reverse CONJ_TAC >- rw [Abbr ‘vs1’, LIST_TO_SET_TAKE] \\ + qunabbrev_tac ‘ys1’ \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘RANK r’ >> rw [DISJOINT_RANK_RNEWS'] \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘ROW 0’ \\ - CONJ_TAC >- rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ - MATCH_MP_TAC ROW_SUBSET_RANK >> art []) >> DISCH_TAC \\ - Know ‘DISJOINT (set ys2) (set zs2)’ + Q.EXISTS_TAC ‘RANK r’ >> simp [DISJOINT_RANK_RNEWS']) >> DISCH_TAC \\ + Know ‘DISJOINT (set vs2) (set ys2)’ >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set vs’ \\ - reverse CONJ_TAC >- rw [Abbr ‘ys2’, LIST_TO_SET_TAKE] \\ - qunabbrev_tac ‘zs2’ \\ + reverse CONJ_TAC >- rw [Abbr ‘vs2’, LIST_TO_SET_TAKE] \\ + qunabbrev_tac ‘ys2’ \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘RANK r’ >> rw [DISJOINT_RANK_RNEWS'] \\ - MATCH_MP_TAC SUBSET_TRANS \\ - Q.EXISTS_TAC ‘ROW 0’ \\ - CONJ_TAC >- rw [Abbr ‘vs’, RNEWS_SUBSET_ROW] \\ - MATCH_MP_TAC ROW_SUBSET_RANK >> art []) >> DISCH_TAC \\ + Q.EXISTS_TAC ‘RANK r’ >> simp [DISJOINT_RANK_RNEWS']) >> DISCH_TAC \\ qabbrev_tac ‘t1 = VAR (y j1) @* args j1’ \\ qabbrev_tac ‘t2 = VAR (y j2) @* args j2’ \\ (* applying for principle_hnf_tpm_reduce *) - Know ‘principle_hnf (LAMl ys1 t1 @* MAP VAR zs1) = tpm (ZIP (ys1,zs1)) t1’ + Know ‘principle_hnf (LAMl vs1 t1 @* MAP VAR ys1) = tpm (ZIP (vs1,ys1)) t1’ >- (‘hnf t1’ by rw [Abbr ‘t1’, hnf_appstar] \\ MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ MATCH_MP_TAC subterm_disjoint_lemma \\ @@ -7097,7 +7193,7 @@ Proof MATCH_MP_TAC SUBSET_TRANS \\ Q.EXISTS_TAC ‘Z’ >> art [] \\ rw [Abbr ‘t1’, FV_appstar]) >> Rewr' \\ - Know ‘principle_hnf (LAMl ys2 t2 @* MAP VAR zs2) = tpm (ZIP (ys2,zs2)) t2’ + Know ‘principle_hnf (LAMl vs2 t2 @* MAP VAR ys2) = tpm (ZIP (vs2,ys2)) t2’ >- (‘hnf t2’ by rw [Abbr ‘t2’, hnf_appstar] \\ MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ MATCH_MP_TAC subterm_disjoint_lemma \\ @@ -7108,16 +7204,154 @@ Proof simp [Abbr ‘t1’, Abbr ‘t2’, tpm_appstar] >> STRIP_TAC \\ Know ‘LENGTH (l j1) = LENGTH (l j2)’ >- (simp [] \\ - ‘n j1 <= n_max /\ n j2 <= n_max’ by rw [] >> simp []) >> DISCH_TAC \\ - reverse CONJ_TAC >- simp [Abbr ‘Ns’, Abbr ‘tl’] \\ + ‘n j1 <= n_max /\ n j2 <= n_max’ by rw [] \\ + simp []) >> DISCH_TAC \\ + reverse CONJ_TAC + >- (simp [Abbr ‘Ns’, Abbr ‘tl’, Abbr ‘d_max'’] \\ + ‘f j1 < k /\ f j2 < k’ by rw [] >> simp []) \\ ‘b j1 = EL (j j1) xs /\ b j2 = EL (j j2) xs’ by rw [] \\ NTAC 2 POP_ORW \\ Suff ‘j j1 = j j2’ >- Rewr \\ simp [Abbr ‘j’, Abbr ‘args'’, Abbr ‘args2’] \\ ‘n j1 <= n_max /\ n j2 <= n_max’ by rw [] \\ - simp []) \\ - (* instantiating the key substitution assumption with q <> [] *) - Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ (MP_TAC o Q.SPEC ‘q’) >> art [] \\ + ‘f j1 < k /\ f j2 < k’ by rw [] \\ + simp [Abbr ‘d_max'’] \\ + Suff ‘f j1 = f j2’ >- rw [] \\ + (* NOTE: current situation: + + |<--------- vs (n_max) --------->| + |<--- vs1 ----->|<---- vs1'----->| y j1 ---+ + |<------ vs2 ------->|<--vs2'--->| y j2 ---|--+ + ----------------------------------------------------|--|---- + |<--- ys1 ----->|------ys1'----->| y' <-----+ | + |<------ ys2 ------->|<--ys2'--->| y' <--------+ + + lswapstr (ZIP (vs, ys)) (y j1) = + lswapstr (ZIP (vs1,ys1)) (y j1) = + lswapstr (ZIP (vs2,ys2)) (y j2) = + lswapstr (ZIP (vs, ys)) (y j2) ==> y j1 = y j2 + + P (f j1) = VAR (y j1) ISUB ss = + VAR (y j2) ISUB ss = P (f j2) + ==> permutator (d_max + f j1) = permutator (d_max + f j2) + ==> d_max + f j1 = d_max + f j2 ==> f j1 = f j2 + *) + PRINT_TAC "stage work on subtree_equiv_lemma: L7252" \\ + Suff ‘y j1 = y j2’ + >- (DISCH_TAC \\ + Know ‘VAR (y j1) ISUB ss = VAR (y j2) ISUB ss’ + >- POP_ASSUM (REWRITE_TAC o wrap) \\ + POP_ASSUM K_TAC \\ + simp [Abbr ‘P’]) (* permutator_11 is used here *) \\ + qabbrev_tac ‘vs1' = DROP (n j1) vs’ \\ + qabbrev_tac ‘vs2' = DROP (n j2) vs’ \\ + Know ‘ys1 <<= ys’ + >- (qunabbrevl_tac [‘ys1’, ‘ys’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n1'’ STRIP_ASSUME_TAC) \\ + Know ‘n1' = n j1’ + >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ + simp [LENGTH_TAKE]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘n1' <= n_max’ MP_TAC \\ + Q.PAT_X_ASSUM ‘ys1 = TAKE n1' ys’ + (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> rpt STRIP_TAC \\ + qabbrev_tac ‘ys1' = DROP (n j1) ys’ \\ + ‘vs1 ++ vs1' = vs /\ ys1 ++ ys1' = ys’ by METIS_TAC [TAKE_DROP] \\ + Know ‘ys2 <<= ys’ + >- (qunabbrevl_tac [‘ys2’, ‘ys’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n2'’ STRIP_ASSUME_TAC) \\ + Know ‘n2' = n j2’ + >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ + simp [LENGTH_TAKE]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘n2' <= n_max’ MP_TAC \\ + Q.PAT_X_ASSUM ‘ys2 = TAKE n2' ys’ + (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> rpt STRIP_TAC \\ + qabbrev_tac ‘ys2' = DROP (n j2) ys’ \\ + ‘vs2 ++ vs2' = vs /\ ys2 ++ ys2' = ys’ by METIS_TAC [TAKE_DROP] \\ + qabbrev_tac ‘pm1 = ZIP (vs1,ys1)’ \\ + qabbrev_tac ‘pm2 = ZIP (vs2,ys2)’ \\ + Suff ‘lswapstr pm1 (y j1) = lswapstr pm (y j1) /\ + lswapstr pm2 (y j2) = lswapstr pm (y j2)’ + >- (STRIP_TAC \\ + Q.PAT_X_ASSUM ‘lswapstr pm1 (y j1) = lswapstr pm2 (y j2)’ MP_TAC \\ + simp []) \\ + Q.PAT_X_ASSUM ‘lswapstr pm1 (y j1) = lswapstr pm2 (y j2)’ K_TAC \\ + CONJ_TAC >| (* 2 subgoals *) + [ (* goal 1 (of 2) *) + ‘LENGTH vs1' = LENGTH ys1'’ by rw [Abbr ‘vs1'’, Abbr ‘ys1'’] \\ + Know ‘pm = pm1 ++ ZIP (vs1',ys1')’ + >- (simp [Abbr ‘pm’, Abbr ‘pm1’] \\ + ‘LENGTH vs1 = LENGTH ys1’ by rw [Abbr ‘vs1'’] \\ + simp [ZIP_APPEND]) >> Rewr' \\ + simp [lswapstr_append, Once EQ_SYM_EQ] \\ + MATCH_MP_TAC lswapstr_unchanged' >> simp [MAP_ZIP] \\ + reverse CONJ_TAC (* easy goal first *) + >- (‘y j1 IN X UNION RANK r’ by METIS_TAC [SUBSET_DEF] \\ + Suff ‘DISJOINT (X UNION RANK r) (set ys1')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set ys’ \\ + reverse CONJ_TAC >- simp [Abbr ‘ys1'’, LIST_TO_SET_DROP] \\ + simp [DISJOINT_UNION', Once DISJOINT_SYM] \\ + simp [Abbr ‘ys’, Once DISJOINT_SYM, DISJOINT_RNEWS_RANK]) \\ + (* current goal: ~MEM (y j1) vs1' + + M0 i = LAMl (TAKE (n i) vs) (VAR (y i) @* args i) + Abbrev (M1 = (\i. principle_hnf (M0 i @* MAP VAR vs))) + M1 i = VAR (y i) @* args i @* DROP (n i) (MAP VAR vs) + + It seems that (y i) at most uses (TAKE (n i) vs). + *) + ‘y j1 IN Y UNION set vs1’ by rw [Abbr ‘vs1’] \\ + Suff ‘DISJOINT (Y UNION set vs1) (set vs1')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + REWRITE_TAC [DISJOINT_UNION] \\ + reverse CONJ_TAC (* easy goal first *) + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs1 ++ vs1' = vs’ (REWRITE_TAC o wrap o SYM) \\ + simp [ALL_DISTINCT_APPEND']) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs’ >> simp [Once DISJOINT_SYM] \\ + simp [Abbr ‘vs1'’, LIST_TO_SET_DROP], + (* goal 2 (of 2) *) + ‘LENGTH vs2' = LENGTH ys2'’ by rw [Abbr ‘vs2'’, Abbr ‘ys2'’] \\ + Know ‘pm = pm2 ++ ZIP (vs2',ys2')’ + >- (simp [Abbr ‘pm’, Abbr ‘pm2’] \\ + ‘LENGTH vs2 = LENGTH ys2’ by rw [Abbr ‘vs2'’] \\ + simp [ZIP_APPEND]) >> Rewr' \\ + simp [lswapstr_append, Once EQ_SYM_EQ] \\ + MATCH_MP_TAC lswapstr_unchanged' >> simp [MAP_ZIP] \\ + reverse CONJ_TAC (* easy goal first *) + >- (‘y j2 IN X UNION RANK r’ by METIS_TAC [SUBSET_DEF] \\ + Suff ‘DISJOINT (X UNION RANK r) (set ys2')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set ys’ \\ + reverse CONJ_TAC >- simp [Abbr ‘ys2'’, LIST_TO_SET_DROP] \\ + simp [DISJOINT_UNION', Once DISJOINT_SYM] \\ + simp [Abbr ‘ys’, Once DISJOINT_SYM, DISJOINT_RNEWS_RANK]) \\ + ‘y j2 IN Y UNION set vs2’ by rw [Abbr ‘vs2’] \\ + Suff ‘DISJOINT (Y UNION set vs2) (set vs2')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + REWRITE_TAC [DISJOINT_UNION] \\ + reverse CONJ_TAC (* easy goal first *) + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs2 ++ vs2' = vs’ (REWRITE_TAC o wrap o SYM) \\ + simp [ALL_DISTINCT_APPEND']) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs’ >> simp [Once DISJOINT_SYM] \\ + simp [Abbr ‘vs2'’, LIST_TO_SET_DROP] ]) \\ + (* stage work, instantiating the key substitution assumption with q <> [] *) + Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ drule >> art [] \\ DISCH_TAC \\ (* NOTE: ‘solvable (subterm' X (M i) q r)’ only holds when ‘q <<= FRONT p’. The case that ‘unsolvable (subterm' X (M j1/j2) q r)’ (p = q) must be @@ -7170,7 +7404,6 @@ Proof rename1 ‘(\(N,r). NONE) z = SOME T’ \\ Cases_on ‘z’ >> FULL_SIMP_TAC std_ss []) \\ (* stage work, now applying BT_subterm_thm on ‘M j1’ *) - PRINT_TAC "[subterm_equiv_lemma] stage work" \\ MP_TAC (Q.SPECL [‘q’, ‘X’, ‘M (j1 :num)’, ‘r’] BT_subterm_thm) \\ simp [] >> STRIP_TAC (* this asserts ‘x’ *) \\ NTAC 3 (Cases_on ‘x’ >> fs []) \\ @@ -7239,8 +7472,8 @@ Proof (* some properties needed by the next "solvable" subgoal *) Know ‘set vs SUBSET X UNION RANK r1’ >- (Suff ‘set vs SUBSET RANK r1’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘RANK r’ >> art [] \\ + simp [Abbr ‘r1’, RANK_MONO]) >> DISCH_TAC \\ Know ‘set ys SUBSET X UNION RANK r1’ >- (Suff ‘set ys SUBSET RANK r1’ >- SET_TAC [] \\ qunabbrev_tac ‘ys’ \\ @@ -7274,15 +7507,23 @@ Proof >- (ASM_SIMP_TAC std_ss [] \\ CONJ_TAC >| (* 2 subgoals *) [ (* goal 1 (of 2) *) - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0’, ‘r1’, ‘d_max’] \\ - simp [principle_hnf_tpm'], + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r1’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + rpt STRIP_TAC \\ + Know ‘y i IN X UNION RANK r’ >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK r1’ >- SET_TAC [] \\ + simp [Abbr ‘r1’, RANK_MONO], (* goal 2 (of 2) *) - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0'’, ‘r1’, ‘d_max’] \\ - simp [principle_hnf_tpm'] ]) >> STRIP_TAC \\ + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r1’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + rpt STRIP_TAC \\ + Know ‘y i IN X UNION RANK r’ >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK r1’ >- SET_TAC [] \\ + simp [Abbr ‘r1’, RANK_MONO] ]) >> STRIP_TAC \\ (* stage work *) - PRINT_TAC "[subterm_equiv_lemma] stage work" \\ + PRINT_TAC "stage work on subtree_equiv_lemma: L7539" \\ Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ ASSUME_TAC \\ (* applying BT_subterm_thm on ‘H j1’ *) MP_TAC (Q.SPECL [‘q’, ‘X’, ‘H (j1 :num)’, ‘r’] BT_subterm_thm) \\ @@ -7322,6 +7563,7 @@ Proof Q_TAC (HNF_TAC (“W0' :term”, “vs4 :string list”, “y4' :string”, “Ns4 :term list”)) ‘W1'’ \\ Q.PAT_X_ASSUM ‘DISJOINT (set vs4) (FV W0')’ K_TAC \\ + (* stage work *) Know ‘TAKE (LAMl_size W0) vs3 = vs3 /\ TAKE (LAMl_size W0') vs4 = vs4’ >- simp [] \\ DISCH_THEN (rfs o CONJUNCTS) \\ @@ -7334,8 +7576,8 @@ Proof (* properties of W0 *) ‘LAMl_size W0 = n3 /\ hnf_children_size W0 = m3 /\ hnf_headvar W1 = y3’ by rw [] \\ - Q.PAT_X_ASSUM ‘W0 = _’ (ASSUME_TAC o SYM) \\ - Q.PAT_X_ASSUM ‘W1 = _’ (ASSUME_TAC o SYM) \\ + Q.PAT_X_ASSUM ‘W0 = _’ (ASSUME_TAC o SYM) \\ + Q.PAT_X_ASSUM ‘W1 = _’ (ASSUME_TAC o SYM) \\ (* properties of W0' *) ‘LAMl_size W0' = n4 /\ hnf_children_size W0' = m4 /\ hnf_headvar W1' = y4’ by rw [] \\ @@ -7377,7 +7619,9 @@ Proof Suff ‘DISJOINT (set vs1) (set vs)’ >- (rw [DISJOINT_ALT] \\ FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) \\ - qunabbrevl_tac [‘vs1’, ‘vs’] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs1’, ‘vs0’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’], (* goal 2 (of 2) *) POP_ASSUM MP_TAC \\ @@ -7397,7 +7641,9 @@ Proof Suff ‘DISJOINT (set vs2) (set vs)’ >- (rw [DISJOINT_ALT] \\ FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) \\ - qunabbrevl_tac [‘vs2’, ‘vs’] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs2’, ‘vs0’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’], (* goal 2 (of 2) *) POP_ASSUM MP_TAC \\ @@ -7443,24 +7689,18 @@ Proof Q.PAT_X_ASSUM ‘LAMl vs4 _ = W0'’ MP_TAC \\ Q.PAT_X_ASSUM ‘_ = W1'’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ - gs []) >> STRIP_TAC \\ - simp [] \\ - (* last goal: m4 + n1 = m3 + n2 *) - qunabbrevl_tac [‘m3’, ‘m4’] \\ + NTAC 2 (POP_ASSUM (fs o wrap))) >> STRIP_TAC \\ + simp [Abbr ‘m3’, Abbr ‘m4’] \\ NTAC 2 (POP_ASSUM (REWRITE_TAC o wrap o SYM)) \\ Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ - simp [Abbr ‘Ns1''’, Abbr ‘Ns2''’] \\ - simp [Abbr ‘Ns1'’, Abbr ‘Ns2'’]) \\ + simp [Abbr ‘Ns1'’, Abbr ‘Ns2'’, Abbr ‘Ns1''’, Abbr ‘Ns2''’]) \\ (* hard case *) - PRINT_TAC "[subterm_equiv_lemma] stage work" \\ - Know ‘VAR y1' ISUB ss = P’ - >- (LAST_X_ASSUM MATCH_MP_TAC \\ - POP_ASSUM MP_TAC >> simp []) >> Rewr' \\ - (* preparing for hreduce_permutator_shared - - NOTE: LENGTH Ns1'' = LENGTH Ns1' = LENGTH Ns1 = m1 <= d_max - *) + PRINT_TAC "stage work on subtree_equiv_lemma: L7712" \\ + POP_ASSUM MP_TAC >> simp [] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘j3’ STRIP_ASSUME_TAC) \\ + ‘(LEAST j. y j = y1') = f j3’ by rw [] >> POP_ORW \\ + (* preparing for hreduce_permutator_shared *) ‘LENGTH Ns1'' = m1 /\ LENGTH Ns2'' = m2’ by simp [Abbr ‘Ns1''’, Abbr ‘Ns2''’, Abbr ‘Ns1'’, Abbr ‘Ns2'’] \\ qabbrev_tac ‘X' = BIGUNION (IMAGE FV (set Ns1'')) UNION @@ -7472,11 +7712,11 @@ Proof It will be shown that SUC d_max + n1 - m1 = n3 = n4. Depending on the independent n1 and m1, either SUC d_max <= n3 or n3 <= SUC d_max. - Thus ‘MAX n3 (SUC d_max)’ is the suitable length to be used here. + Thus ‘MAX n3 (SUC (d_max' j3))’ is the suitable length to be used here. *) - qabbrev_tac ‘d' = MAX n3 (SUC d_max)’ \\ + qabbrev_tac ‘d' = MAX n3 (SUC (d_max' j3))’ \\ Q_TAC (NEWS_TAC (“L :string list”, “d' :num”)) ‘X'’ \\ - ‘SUC d_max <= LENGTH L /\ n3 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ + ‘d_max' j3 < LENGTH L /\ n3 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ Know ‘DISJOINT (set L) (set vs1) /\ DISJOINT (set L) (set vs2) /\ DISJOINT (set L) (set vs3) /\ @@ -7488,17 +7728,19 @@ Proof qunabbrev_tac ‘X'’ \\ DISCH_THEN (STRIP_ASSUME_TAC o REWRITE_RULE [DISJOINT_UNION']) \\ STRIP_TAC (* W -h->* ... /\ W' -h->* ... *) \\ + ‘m1 <= d_max' j3 /\ m2 <= d_max' j3’ by simp [Abbr ‘d_max'’] \\ (* applying hreduce_permutator_shared *) - MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max’, ‘L’] hreduce_permutator_shared) \\ - simp [] \\ + MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max + f (j3 :num)’, ‘L’] + hreduce_permutator_shared) >> simp [] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘zs1’ (Q.X_CHOOSE_THEN ‘z1’ STRIP_ASSUME_TAC)) \\ (* applying hreduce_permutator_shared again *) - MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max’, ‘L’] hreduce_permutator_shared) \\ - simp [] \\ + MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max + f (j3 :num)’, ‘L’] + hreduce_permutator_shared) >> simp [] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘zs2’ (Q.X_CHOOSE_THEN ‘z2’ STRIP_ASSUME_TAC)) \\ - Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ MP_TAC \\ - Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ MP_TAC \\ - qabbrev_tac ‘h = LAST L’ (* the new shared head variable *) \\ + qabbrev_tac ‘P' = P (f j3)’ \\ + Q.PAT_X_ASSUM ‘P' @* Ns1'' -h->* _’ MP_TAC \\ + Q.PAT_X_ASSUM ‘P' @* Ns2'' -h->* _’ MP_TAC \\ + qabbrev_tac ‘h = LAST L’ (* new shared head variable *) \\ qabbrev_tac ‘L' = FRONT L’ \\ ‘L <> []’ by rw [GSYM LENGTH_NON_NIL] \\ NTAC 2 (Q.PAT_X_ASSUM ‘IS_SUFFIX L _’ MP_TAC) \\ @@ -7515,13 +7757,13 @@ Proof by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] \\ POP_ORW \\ simp [IS_SUFFIX, Abbr ‘xs1’, Abbr ‘xs2’]) >> STRIP_TAC \\ - Know ‘LAMl vs1 (P @* Ns1'') -h->* + Know ‘LAMl vs1 (P' @* Ns1'') -h->* LAMl vs1 (LAMl zs1 (LAM h (VAR h @* Ns1'' @* MAP VAR zs1))) /\ - LAMl vs2 (P @* Ns2'') -h->* + LAMl vs2 (P' @* Ns2'') -h->* LAMl vs2 (LAMl zs2 (LAM h (VAR h @* Ns2'' @* MAP VAR zs2)))’ >- simp [hreduce_LAMl] \\ - Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ K_TAC \\ - Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ K_TAC \\ + Q.PAT_X_ASSUM ‘P' @* Ns1'' -h->* _’ K_TAC \\ + Q.PAT_X_ASSUM ‘P' @* Ns2'' -h->* _’ K_TAC \\ REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND] \\ qabbrev_tac ‘Ns1x = Ns1'' ++ MAP VAR zs1’ \\ qabbrev_tac ‘Ns2x = Ns2'' ++ MAP VAR zs2’ \\ @@ -7544,19 +7786,18 @@ Proof Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ simp []) >> STRIP_TAC \\ - (* n3 = LENGTH zs1' = 1 + LENGTH (vs1 ++ zs1) = 1 + d_max + n1 - m1 - - NOTE: ‘LENGTH L = SUC d_max’ - *) - Know ‘SUC d_max + n1 - m1 = n3 /\ - SUC d_max + n1 - m1 = n4’ + (* n3 = LENGTH zs1' = 1 + LENGTH (vs1 ++ zs1) = 1 + d_max + n1 - m1 *) + Know ‘SUC (d_max' j3) + n1 - m1 = n3 /\ + SUC (d_max' j3) + n2 - m2 = n4’ >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ simp [Abbr ‘zs1'’, Abbr ‘zs2'’]) >> STRIP_TAC \\ - ‘n4 = n3’ by PROVE_TAC [] \\ - ‘vs4 = vs3’ by simp [Abbr ‘vs4’, Abbr ‘vs3’] >> simp [] \\ + Know ‘n4 = n3’ + >- (NTAC 2 (POP_ASSUM (REWRITE_TAC o wrap o SYM)) \\ + simp []) >> DISCH_TAC \\ + ‘vs4 = vs3’ by simp [Abbr ‘vs4’, Abbr ‘vs3’] \\ simp [head_equivalent_def] \\ Know ‘m4 = m3’ >- (Q.PAT_X_ASSUM ‘hnf_children_size W0 = m3’ (REWRITE_TAC o wrap o SYM) \\ @@ -7564,7 +7805,7 @@ Proof Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ simp [Abbr ‘Ns1x’, Abbr ‘Ns2x’]) >> DISCH_TAC \\ - simp [] \\ + simp [] (* it remains to prove ‘y3 = y4’ *) \\ (* NOTE: Now we know: 1. LAMl zs1' (VAR h @* Ns1x) = W0 @@ -7601,9 +7842,10 @@ Proof zs2' = |<--- vs2 ------>|<-- zs2 ---->|h| (ROW 0 & r1) vs3 = |<---(vs2)----- vs4 ---(xs2)---->| (ROW r1) (n4 = 1 + d_max + n2 - m2 > n2) + + now applying RNEWS_prefix first: *) - PRINT_TAC "[subterm_equiv_lemma] stage work" \\ - (* applying RNEWS_prefix *) + PRINT_TAC "stage work on subtree_equiv_lemma: L7861" \\ Know ‘vs1 <<= vs3’ >- (qunabbrevl_tac [‘vs1’, ‘vs3’] \\ MATCH_MP_TAC RNEWS_prefix >> simp []) \\ @@ -7641,9 +7883,8 @@ Proof MATCH_MP_TAC DISJOINT_SUBSET' \\ Q.EXISTS_TAC ‘set L’ >> art [] \\ MATCH_MP_TAC LIST_TO_SET_SUFFIX >> art []) >> DISCH_TAC \\ - Know ‘LENGTH xs1 = LENGTH ys1 /\ LENGTH xs2 = LENGTH ys2’ - >- simp [Abbr ‘xs1’, Abbr ‘ys1’, Abbr ‘xs2’, Abbr ‘ys2’] \\ - STRIP_TAC \\ + ‘LENGTH xs1 = LENGTH ys1 /\ LENGTH xs2 = LENGTH ys2’ + by simp [Abbr ‘xs1’, Abbr ‘ys1’, Abbr ‘xs2’, Abbr ‘ys2’] \\ ‘vs1 ++ ys1 = vs3 /\ vs2 ++ ys2 = vs3’ by METIS_TAC [TAKE_DROP] \\ (* applying hreduce_BETA_extended: W0 @* MAP VAR vs3 @@ -7671,7 +7912,9 @@ Proof MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs3’ \\ simp [Abbr ‘ys1’, Abbr ‘ys2’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘vs’, ‘vs3’] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘vs3’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> STRIP_TAC \\ Know ‘DISJOINT (set ys) (set ys1) /\ DISJOINT (set ys) (set ys2)’ @@ -7720,22 +7963,13 @@ Proof reverse CONJ_TAC >- (MP_TAC (Q.SPECL [‘ss’, ‘EL i Ns1'’] FV_ISUB_upperbound) \\ simp [EL_MAP, Abbr ‘Ns1'’]) \\ - (* Ns1' is tpm (REVERSE pm) of Ns1 - pm = ZIP (vs,ys), vs is in ROW 0, ys s in ROW r - vsx (part of vs4) is in ROW r1 > r - - The key is to prove DISJOINT (set vsx) (FV (EL i Ns1)). - *) + (* The key is to prove DISJOINT (set vsx) (FV (EL i Ns1)) *) POP_ASSUM MP_TAC \\ simp [Abbr ‘Ns1'’, EL_listpm, Abbr ‘pm’, REVERSE_ZIP] >> DISCH_TAC \\ (* applying FV_tpm_disjoint *) MATCH_MP_TAC FV_tpm_disjoint \\ simp [ALL_DISTINCT_REVERSE] \\ - (* goal: DISJOINT (set ys1) (FV (EL i Ns1)), given: - - principle_hnf (N0 @* MAP VAR vs1) = VAR y1 @* Ns1 - FV N0 SUBSET FV N SUBSET X UNION RANK r1 - *) + (* goal: DISJOINT (set ys1) (FV (EL i Ns1)) *) Know ‘FV N0 SUBSET X UNION RANK r1’ >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV N’ >> art [] \\ qunabbrev_tac ‘N0’ \\ @@ -7815,7 +8049,6 @@ Proof qunabbrev_tac ‘vs3’ \\ MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> DISCH_TAC \\ (* stage work *) - PRINT_TAC "[subterm_equiv_lemma] stage work" \\ qabbrev_tac ‘pm1 = ZIP (xs1,ys1)’ \\ qabbrev_tac ‘pm2 = ZIP (xs2,ys2)’ \\ ‘W0 @* MAP VAR vs3 -h->* tpm pm1 (VAR h @* Ns1x) /\ @@ -7857,9 +8090,10 @@ Proof qabbrev_tac ‘pm3 = ZIP (ls,vs3)’ \\ (* applying IS_SUFFIX_IMP_LASTN *) Know ‘DROP n1 ls = xs1 /\ DROP n2 ls = xs2’ - >- (‘xs1 = LASTN (LENGTH xs1) L’ by PROVE_TAC [IS_SUFFIX_IMP_LASTN] \\ + >- (PRINT_TAC "stage work on subtree_equiv_lemma: L8106" \\ + ‘xs1 = LASTN (LENGTH xs1) L’ by simp [IS_SUFFIX_IMP_LASTN] \\ POP_ORW \\ - ‘xs2 = LASTN (LENGTH xs2) L’ by PROVE_TAC [IS_SUFFIX_IMP_LASTN] \\ + ‘xs2 = LASTN (LENGTH xs2) L’ by simp [IS_SUFFIX_IMP_LASTN] \\ POP_ORW \\ MP_TAC (ISPECL [“n1 :num”, “ls :string list”] DROP_LASTN) \\ MP_TAC (ISPECL [“n2 :num”, “ls :string list”] DROP_LASTN) \\ @@ -7870,11 +8104,8 @@ Proof Know ‘LASTN (n3 - n2) (LASTN n3 L) = LASTN (n3 - n2) L’ >- (irule LASTN_LASTN >> simp []) >> Rewr' \\ Suff ‘LENGTH ys1 = n3 - n1 /\ LENGTH ys2 = n3 - n2’ >- Rewr \\ - Know ‘LENGTH (vs1 ++ ys1) = LENGTH vs3 /\ - LENGTH (vs2 ++ ys2) = LENGTH vs3’ - >- (Q.PAT_X_ASSUM ‘vs1 ++ ys1 = vs3’ (REWRITE_TAC o wrap) \\ - Q.PAT_X_ASSUM ‘vs2 ++ ys2 = vs3’ (REWRITE_TAC o wrap)) \\ - ASM_SIMP_TAC list_ss [LENGTH_APPEND]) >> STRIP_TAC \\ + simp [Abbr ‘ys1’, Abbr ‘ys2’, LENGTH_DROP]) >> STRIP_TAC \\ + PRINT_TAC "stage work on subtree_equiv_lemma: L8121" \\ (* preparing for lswapstr_unchanged' *) qabbrev_tac ‘xs1' = TAKE n1 ls’ \\ qabbrev_tac ‘xs2' = TAKE n2 ls’ \\ @@ -7960,27 +8191,25 @@ Proof subtree_equiv X (apply pi M) (apply pi N) p r ==> subtree_equiv' X M N p r *) - >> qx_genl_tac [‘t1’, ‘t2’] (* these names are soon replaced by M' *) - >> STRIP_TAC + >> PRINT_TAC "stage work on subtree_equiv_lemma: L8207" + >> rpt GEN_TAC >> STRIP_TAC >> POP_ASSUM MP_TAC >> ONCE_REWRITE_TAC [MONO_NOT_EQ] (* NOTE: The antecedent “~subtree_equiv' X t1 t2 q r” makes sure that ‘n1 + m2 <> n1 + m1’ is always assumed (instead of ‘y1 <> y2’). And - the goal is to prove ‘y3 <> y4 \/ n3 + m3 <> n4 + m4’ - - The original proof assumes q = p, but the proof also work for q <<= p. + the goal is to prove ‘y3 <> y4 \/ n3 + m3 <> n4 + m4’. *) - >> NTAC 2 (POP_ASSUM MP_TAC) + >> NTAC 2 (Q.PAT_X_ASSUM ‘MEM _ Ms’ MP_TAC) >> simp [MEM_EL] >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j1’ STRIP_ASSUME_TAC) >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j2’ STRIP_ASSUME_TAC) >> Q.PAT_X_ASSUM ‘_ = M j1’ (REWRITE_TAC o wrap) >> Q.PAT_X_ASSUM ‘_ = M j2’ (REWRITE_TAC o wrap) >> qabbrev_tac ‘M' = \i. apply pi (M i)’ - (* goal: ~subtree_equiv X (M j1) (M j2) p r ==> - ~subtree_equiv X (M' j1) (M' j2) p r + (* real goal: ~subtree_equiv X (M j1) (M j2) p r ==> + ~subtree_equiv X (M' j1) (M' j2) p r *) - >> simp [subtree_equiv_def, subtree_equiv'_def] + >> simp [subtree_equiv_def] >> Know ‘BT' X (M' j1) r = BT' X (principle_hnf (M' j1)) r’ >- (ONCE_REWRITE_TAC [EQ_SYM_EQ] \\ MATCH_MP_TAC BT_of_principle_hnf \\ @@ -7993,37 +8222,240 @@ Proof simp [Abbr ‘M'’] \\ METIS_TAC [lameq_solvable_cong]) >> Rewr' - >> simp [Abbr ‘M'’] (* now H is involved *) + >> simp [Abbr ‘M'’] (* now H is involved instead of ‘apply pi ...’ *) + (* special case: q = [] *) + >> Cases_on ‘q = []’ + >- (Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ K_TAC \\ + POP_ORW >> simp [BT_ltree_el_NIL] \\ + Know ‘!i. principle_hnf (H i) = H i’ + >- (rw [Abbr ‘H’] >> MATCH_MP_TAC principle_hnf_reduce \\ + rw [hnf_appstar]) >> Rewr' \\ + Q.PAT_X_ASSUM ‘!i. solvable (H i)’ K_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> FV (H i) SUBSET X UNION RANK r’ K_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> d_max <= LENGTH (hnf_children (H i))’ K_TAC \\ + simp [Abbr ‘H’, GSYM appstar_APPEND, hnf_head_appstar] \\ + simp [head_equivalent_def] \\ + qabbrev_tac ‘vs1 = TAKE (n j1) vs’ \\ + qabbrev_tac ‘vs2 = TAKE (n j2) vs’ \\ + ‘ALL_DISTINCT vs1 /\ ALL_DISTINCT vs2’ + by simp [Abbr ‘vs1’, Abbr ‘vs2’, ALL_DISTINCT_TAKE] \\ + ‘LENGTH vs1 = n j1’ + by (qunabbrev_tac ‘vs1’ \\ + MATCH_MP_TAC LENGTH_TAKE >> art [] \\ + FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ + ‘LENGTH vs2 = n j2’ + by (qunabbrev_tac ‘vs2’ \\ + MATCH_MP_TAC LENGTH_TAKE >> art [] \\ + FIRST_X_ASSUM MATCH_MP_TAC >> art []) \\ + Q_TAC (RNEWS_TAC (“ys1 :string list”, “r :num”, + “(n :num -> num) j1”)) ‘X’ \\ + Q_TAC (RNEWS_TAC (“ys2 :string list”, “r :num”, + “(n :num -> num) j2”)) ‘X’ \\ + Know ‘DISJOINT (set vs1) (set ys1)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs’ \\ + reverse CONJ_TAC >- rw [Abbr ‘vs1’, LIST_TO_SET_TAKE] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘ys1’] \\ + MATCH_MP_TAC DISJOINT_RNEWS >> simp []) >> DISCH_TAC \\ + Know ‘DISJOINT (set vs2) (set ys2)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs’ \\ + reverse CONJ_TAC >- rw [Abbr ‘vs2’, LIST_TO_SET_TAKE] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘ys2’] \\ + MATCH_MP_TAC DISJOINT_RNEWS >> simp []) >> DISCH_TAC \\ + qabbrev_tac ‘t1 = VAR (y j1) @* args j1’ \\ + qabbrev_tac ‘t2 = VAR (y j2) @* args j2’ \\ + (* applying for principle_hnf_tpm_reduce *) + Know ‘principle_hnf (LAMl vs1 t1 @* MAP VAR ys1) = tpm (ZIP (vs1,ys1)) t1’ + >- (‘hnf t1’ by rw [Abbr ‘t1’, hnf_appstar] \\ + MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ + MATCH_MP_TAC subterm_disjoint_lemma \\ + qexistsl_tac [‘X’, ‘r’, ‘n j1’] >> simp [] \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘Z’ >> art [] \\ + rw [Abbr ‘t1’, FV_appstar]) >> Rewr' \\ + Know ‘principle_hnf (LAMl vs2 t2 @* MAP VAR ys2) = tpm (ZIP (vs2,ys2)) t2’ + >- (‘hnf t2’ by rw [Abbr ‘t2’, hnf_appstar] \\ + MATCH_MP_TAC principle_hnf_tpm_reduce' >> art [] \\ + MATCH_MP_TAC subterm_disjoint_lemma \\ + qexistsl_tac [‘X’, ‘r’, ‘n j2’] >> simp [] \\ + MATCH_MP_TAC SUBSET_TRANS \\ + Q.EXISTS_TAC ‘Z’ >> art [] \\ + rw [Abbr ‘t2’, FV_appstar]) >> Rewr' \\ + simp [Abbr ‘t1’, Abbr ‘t2’, tpm_appstar] \\ + PRINT_TAC "stage work on subtree_equiv_lemma: L8303" \\ + qabbrev_tac ‘pm1 = ZIP (vs1,ys1)’ \\ + qabbrev_tac ‘pm2 = ZIP (vs2,ys2)’ \\ + qabbrev_tac ‘vs1' = DROP (n j1) vs’ \\ + qabbrev_tac ‘vs2' = DROP (n j2) vs’ \\ + (* current situation: + |<--------- vs (n_max) --------->| + |<--- vs1 ----->|<---- vs1'----->| y j1 ---+ + |<------ vs2 ------->|<--vs2'--->| y j2 ---|--+ + ----------------------------------------------------|--|---- + |<--- ys1 ----->|------ys1'----->| y' <-----+ | + |<------ ys2 ------->|<--ys2'--->| y' <--------+ + *) + Know ‘ys1 <<= ys’ + >- (qunabbrevl_tac [‘ys1’, ‘ys’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n1'’ STRIP_ASSUME_TAC) \\ + Know ‘n1' = n j1’ + >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ + simp [LENGTH_TAKE]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘n1' <= n_max’ MP_TAC \\ + Q.PAT_X_ASSUM ‘ys1 = TAKE n1' ys’ + (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> NTAC 2 DISCH_TAC \\ + qabbrev_tac ‘ys1' = DROP (n j1) ys’ \\ + ‘vs1 ++ vs1' = vs /\ ys1 ++ ys1' = ys’ by METIS_TAC [TAKE_DROP] \\ + Know ‘ys2 <<= ys’ + >- (qunabbrevl_tac [‘ys2’, ‘ys’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n2'’ STRIP_ASSUME_TAC) \\ + Know ‘n2' = n j2’ + >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ + simp [LENGTH_TAKE]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘n2' <= n_max’ MP_TAC \\ + Q.PAT_X_ASSUM ‘ys2 = TAKE n2' ys’ + (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> NTAC 2 DISCH_TAC \\ + qabbrev_tac ‘ys2' = DROP (n j2) ys’ \\ + ‘vs2 ++ vs2' = vs /\ ys2 ++ ys2' = ys’ by METIS_TAC [TAKE_DROP] \\ + (* stage work *) + Know ‘lswapstr pm1 (y j1) = lswapstr pm (y j1)’ + >- (‘LENGTH vs1' = LENGTH ys1'’ by rw [Abbr ‘vs1'’, Abbr ‘ys1'’] \\ + Know ‘pm = pm1 ++ ZIP (vs1',ys1')’ + >- (simp [Abbr ‘pm’, Abbr ‘pm1’] \\ + ‘LENGTH vs1 = LENGTH ys1’ by rw [Abbr ‘vs1'’] \\ + simp [ZIP_APPEND]) >> Rewr' \\ + simp [lswapstr_append, Once EQ_SYM_EQ] \\ + MATCH_MP_TAC lswapstr_unchanged' >> simp [MAP_ZIP] \\ + reverse CONJ_TAC (* easy goal first *) + >- (‘y j1 IN X UNION RANK r’ by METIS_TAC [SUBSET_DEF] \\ + Suff ‘DISJOINT (X UNION RANK r) (set ys1')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set ys’ \\ + reverse CONJ_TAC >- simp [Abbr ‘ys1'’, LIST_TO_SET_DROP] \\ + simp [DISJOINT_UNION', Once DISJOINT_SYM] \\ + simp [Abbr ‘ys’, Once DISJOINT_SYM, DISJOINT_RNEWS_RANK]) \\ + ‘y j1 IN Y UNION set vs1’ by rw [Abbr ‘vs1’] \\ + Suff ‘DISJOINT (Y UNION set vs1) (set vs1')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + REWRITE_TAC [DISJOINT_UNION] \\ + reverse CONJ_TAC (* easy goal first *) + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs1 ++ vs1' = vs’ (REWRITE_TAC o wrap o SYM) \\ + simp [ALL_DISTINCT_APPEND']) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs’ >> simp [Once DISJOINT_SYM] \\ + simp [Abbr ‘vs1'’, LIST_TO_SET_DROP]) >> Rewr' \\ + Know ‘lswapstr pm2 (y j2) = lswapstr pm (y j2)’ + >- (‘LENGTH vs2' = LENGTH ys2'’ by rw [Abbr ‘vs2'’, Abbr ‘ys2'’] \\ + Know ‘pm = pm2 ++ ZIP (vs2',ys2')’ + >- (simp [Abbr ‘pm’, Abbr ‘pm2’] \\ + ‘LENGTH vs2 = LENGTH ys2’ by rw [Abbr ‘vs2'’] \\ + simp [ZIP_APPEND]) >> Rewr' \\ + simp [lswapstr_append, Once EQ_SYM_EQ] \\ + MATCH_MP_TAC lswapstr_unchanged' >> simp [MAP_ZIP] \\ + reverse CONJ_TAC (* easy goal first *) + >- (‘y j2 IN X UNION RANK r’ by METIS_TAC [SUBSET_DEF] \\ + Suff ‘DISJOINT (X UNION RANK r) (set ys2')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set ys’ \\ + reverse CONJ_TAC >- simp [Abbr ‘ys2'’, LIST_TO_SET_DROP] \\ + simp [DISJOINT_UNION', Once DISJOINT_SYM] \\ + simp [Abbr ‘ys’, Once DISJOINT_SYM, DISJOINT_RNEWS_RANK]) \\ + ‘y j2 IN Y UNION set vs2’ by rw [Abbr ‘vs2’] \\ + Suff ‘DISJOINT (Y UNION set vs2) (set vs2')’ + >- (REWRITE_TAC [DISJOINT_ALT] \\ + DISCH_THEN MATCH_MP_TAC >> art []) \\ + REWRITE_TAC [DISJOINT_UNION] \\ + reverse CONJ_TAC (* easy goal first *) + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs2 ++ vs2' = vs’ (REWRITE_TAC o wrap o SYM) \\ + simp [ALL_DISTINCT_APPEND']) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs’ >> simp [Once DISJOINT_SYM] \\ + simp [Abbr ‘vs2'’, LIST_TO_SET_DROP]) >> Rewr' \\ + simp [] \\ + ‘f j1 < k /\ f j2 < k’ by rw [] \\ + ‘b j1 = EL (j j1) xs /\ b j2 = EL (j j2) xs’ by rw [] \\ + NTAC 2 POP_ORW \\ + Know ‘EL (j j1) xs = EL (j j2) xs <=> j j1 = j j2’ + >- (Q.PAT_X_ASSUM ‘!i. i < k ==> EL (j i) xs = b i /\ _’ K_TAC \\ + MATCH_MP_TAC ALL_DISTINCT_EL_IMP >> art [] \\ + simp [Abbr ‘j’, Abbr ‘args2’, Abbr ‘d_max'’]) >> Rewr' \\ + (* !i. i < k ==> LENGTH (args' i ++ args2 i) <= d_max' i *) + Know ‘LENGTH (args' j1 ++ args2 j1) <= d_max' j1 /\ + LENGTH (args' j2 ++ args2 j2) <= d_max' j2’ + >- simp [] \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> LENGTH (args' i ++ args2 i) <= d_max' i’ K_TAC \\ + simp [Abbr ‘Ns’, Abbr ‘tl’, Abbr ‘d_max'’, Abbr ‘args2’] \\ + (* arithmetic cleanups *) + Know ‘d_max + (k + (n_max + (f j2 + (m j2 + SUC d_max)))) - + (n j2 + SUC (d_max + f j2)) = + d_max + k + n_max + m j2 - n j2’ >- simp [] >> Rewr' \\ + Know ‘d_max + (k + (n_max + (f j1 + (m j1 + SUC d_max)))) - + (n j1 + SUC (d_max + f j1)) = + d_max + k + n_max + m j1 - n j1’ >- simp [] >> Rewr' \\ + Know ‘d_max + k + n_max + m j2 - n j2 = + d_max + k + n_max + m j1 - n j1 <=> m j2 + n j1 = m j1 + n j2’ + >- simp [] >> Rewr' \\ + simp [Abbr ‘j’] \\ + Cases_on ‘y j1 = y j2’ >> simp [] (* only one goal is left *) \\ + Cases_on ‘m j2 + n j1 = m j1 + n j2’ >> simp [] (* 1s left *) \\ + ‘m j1 <= d /\ m j2 <= d’ by rw [] \\ + STRIP_TAC \\ + qabbrev_tac ‘a1 = n_max + m j1’ \\ + qabbrev_tac ‘a2 = n_max + m j2’ \\ + qabbrev_tac ‘b1 = d_max + (f j1 + n j1)’ \\ + qabbrev_tac ‘b2 = d_max + (f j2 + n j2)’ \\ + Know ‘b1 - a1 <> b2 - a2 <=> b1 + a2 <> b2 + a1’ >- simp [] >> Rewr' \\ + qunabbrevl_tac [‘a1’, ‘a2’, ‘b1’, ‘b2’] \\ + Suff ‘f j1 <> f j2’ >- simp [Abbr ‘d_max’] \\ + FIRST_X_ASSUM MATCH_MP_TAC >> art []) (* instantiating the key substitution assumption with q <> [] *) - >> Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ (MP_TAC o Q.SPEC ‘p’) + >> Q.PAT_X_ASSUM ‘!q. q <<= p /\ q <> [] ==> _’ (MP_TAC o Q.SPEC ‘q’) >> simp [] >> DISCH_TAC - (* stage work *) - >> reverse (Cases_on ‘solvable (subterm' X (M j1) p r)’) - >- (Know ‘unsolvable (subterm' X (M j1) p r) <=> - ltree_el (BT' X (M j1) r) p = SOME bot’ + (* some easy cases *) + >> reverse (Cases_on ‘solvable (subterm' X (M j1) q r)’) + >- (Know ‘unsolvable (subterm' X (M j1) q r) <=> + ltree_el (BT' X (M j1) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> rw []) \\ simp [] >> DISCH_THEN K_TAC \\ - reverse (Cases_on ‘solvable (subterm' X (M j2) p r)’) - >- (Know ‘unsolvable (subterm' X (M j2) p r) <=> - ltree_el (BT' X (M j2) r) p = SOME bot’ + reverse (Cases_on ‘solvable (subterm' X (M j2) q r)’) + >- (Know ‘unsolvable (subterm' X (M j2) q r) <=> + ltree_el (BT' X (M j2) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> rw []) \\ - simp [ltree_equiv_imp_ltree_equiv']) \\ - Know ‘unsolvable (subterm' X (H j1) p r)’ + simp []) \\ + Know ‘unsolvable (subterm' X (H j1) q r)’ >- (ASM_SIMP_TAC std_ss [] \\ MATCH_MP_TAC unsolvable_ISUB \\ simp [solvable_tpm]) >> DISCH_TAC \\ - Know ‘unsolvable (subterm' X (H j1) p r) <=> - ltree_el (BT' X (H j1) r) p = SOME bot’ + Know ‘unsolvable (subterm' X (H j1) q r) <=> + ltree_el (BT' X (H j1) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> simp []) \\ simp [] >> DISCH_THEN K_TAC \\ - MP_TAC (Q.SPECL [‘p’, ‘X’, ‘M (j2 :num)’, ‘r’] BT_subterm_thm) \\ + MP_TAC (Q.SPECL [‘q’, ‘X’, ‘M (j2 :num)’, ‘r’] BT_subterm_thm) \\ simp [] >> STRIP_TAC \\ NTAC 3 (Cases_on ‘x’ >> fs []) \\ - qabbrev_tac ‘r2 = r + LENGTH p’ \\ - rename1 ‘subterm X (M j2) p r = SOME (N,r2)’ \\ + qabbrev_tac ‘r2 = r + LENGTH q’ \\ + rename1 ‘subterm X (M j2) q r = SOME (N,r2)’ \\ qabbrev_tac ‘N0 = principle_hnf N’ \\ qabbrev_tac ‘m2 = hnf_children_size N0’ \\ - rename1 ‘ltree_el (BT' X (M j2) r) p = SOME (SOME (vs2,y2),SOME m2)’ \\ + rename1 ‘ltree_el (BT' X (M j2) r) q = SOME (SOME (vs2,y2),SOME m2)’ \\ Q.PAT_X_ASSUM ‘_ = SOME (vs2,y2)’ K_TAC >> gs [] \\ Q.PAT_X_ASSUM ‘_ = r2’ K_TAC \\ Q.PAT_X_ASSUM ‘_ = SOME m2’ K_TAC \\ @@ -8036,7 +8468,7 @@ Proof “y2' :string”, “Ns2 :term list”)) ‘N1’ \\ ‘TAKE (LAMl_size N0) vs2 = vs2’ by rw [] \\ POP_ASSUM (rfs o wrap) \\ - ‘subterm X (H j2) p r <> NONE’ by ASM_SIMP_TAC std_ss [] \\ + ‘subterm X (H j2) q r <> NONE’ by ASM_SIMP_TAC std_ss [] \\ Know ‘IMAGE y (count k) SUBSET X UNION RANK r2’ >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ \\ reverse CONJ_TAC @@ -8048,8 +8480,8 @@ Proof FIRST_X_ASSUM ACCEPT_TAC) >> DISCH_TAC \\ Know ‘set vs SUBSET X UNION RANK r2’ >- (Suff ‘set vs SUBSET RANK r2’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> simp [Abbr ‘r2’]) >> DISCH_TAC \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘RANK r’ >> art [] \\ + simp [Abbr ‘r2’, RANK_MONO]) >> DISCH_TAC \\ Know ‘set ys SUBSET X UNION RANK r2’ >- (Suff ‘set ys SUBSET RANK r2’ >- SET_TAC [] \\ qunabbrev_tac ‘ys’ \\ @@ -8062,48 +8494,54 @@ Proof Know ‘m2 <= d_max’ >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘d’ \\ reverse CONJ_TAC >- rw [Abbr ‘d_max’] \\ - Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j2) p’ \\ - reverse CONJ_TAC >- simp [] \\ + Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j2) q’ \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC subterm_width_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ qunabbrevl_tac [‘m2’, ‘N0’] \\ - ‘N = subterm' X (M j2) p r’ by rw [] >> POP_ORW \\ + ‘N = subterm' X (M j2) q r’ by rw [] >> POP_ORW \\ MATCH_MP_TAC subterm_width_last >> simp []) >> DISCH_TAC \\ - Know ‘solvable (subterm' X (H j2) p r)’ + Know ‘solvable (subterm' X (H j2) q r)’ >- (ASM_SIMP_TAC std_ss [] \\ - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0’, ‘r2’, ‘d_max’] \\ - simp [principle_hnf_tpm']) >> DISCH_TAC \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) p r <> NONE /\ _’ ASSUME_TAC \\ - MP_TAC (Q.SPECL [‘p’, ‘X’, ‘H (j2 :num)’, ‘r’] BT_subterm_thm) \\ + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r2’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + POP_ASSUM MP_TAC >> rw [Abbr ‘m2’] \\ + Know ‘y i IN X UNION RANK r’ >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK r2’ >- SET_TAC [] \\ + rw [Abbr ‘r2’, RANK_MONO]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ ASSUME_TAC \\ + MP_TAC (Q.SPECL [‘q’, ‘X’, ‘H (j2 :num)’, ‘r’] BT_subterm_thm) \\ simp [] >> STRIP_TAC (* this asserts ‘x’ *) \\ NTAC 3 (Cases_on ‘x’ >> fs []) \\ - simp [head_equivalent_def, head_equivalent'_def]) + simp [head_equivalent_def]) (* stage work *) - >> reverse (Cases_on ‘solvable (subterm' X (M j2) p r)’) - >- (Know ‘unsolvable (subterm' X (M j2) p r) <=> - ltree_el (BT' X (M j2) r) p = SOME bot’ + >> reverse (Cases_on ‘solvable (subterm' X (M j2) q r)’) + >- (Know ‘unsolvable (subterm' X (M j2) q r) <=> + ltree_el (BT' X (M j2) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> rw []) \\ simp [] >> DISCH_THEN K_TAC \\ - reverse (Cases_on ‘solvable (subterm' X (M j1) p r)’) - >- (Know ‘unsolvable (subterm' X (M j1) p r) <=> - ltree_el (BT' X (M j1) r) p = SOME bot’ + reverse (Cases_on ‘solvable (subterm' X (M j1) q r)’) + >- (Know ‘unsolvable (subterm' X (M j1) q r) <=> + ltree_el (BT' X (M j1) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> rw []) \\ - simp [ltree_equiv_imp_ltree_equiv']) \\ - Know ‘unsolvable (subterm' X (H j2) p r)’ + simp []) \\ + Know ‘unsolvable (subterm' X (H j2) q r)’ >- (ASM_SIMP_TAC std_ss [] \\ MATCH_MP_TAC unsolvable_ISUB \\ simp [solvable_tpm]) >> DISCH_TAC \\ - Know ‘unsolvable (subterm' X (H j2) p r) <=> - ltree_el (BT' X (H j2) r) p = SOME bot’ + Know ‘unsolvable (subterm' X (H j2) q r) <=> + ltree_el (BT' X (H j2) r) q = SOME bot’ >- (MATCH_MP_TAC BT_ltree_el_of_unsolvables >> simp []) \\ simp [] >> DISCH_THEN K_TAC \\ - MP_TAC (Q.SPECL [‘p’, ‘X’, ‘M (j1 :num)’, ‘r’] BT_subterm_thm) \\ + MP_TAC (Q.SPECL [‘q’, ‘X’, ‘M (j1 :num)’, ‘r’] BT_subterm_thm) \\ simp [] >> STRIP_TAC \\ NTAC 3 (Cases_on ‘x’ >> fs []) \\ - qabbrev_tac ‘r1 = r + LENGTH p’ \\ - rename1 ‘subterm X (M j1) p r = SOME (N,r1)’ \\ + qabbrev_tac ‘r1 = r + LENGTH q’ \\ + rename1 ‘subterm X (M j1) q r = SOME (N,r1)’ \\ qabbrev_tac ‘N0 = principle_hnf N’ \\ qabbrev_tac ‘m1 = hnf_children_size N0’ \\ - rename1 ‘ltree_el (BT' X (M j1) r) p = SOME (SOME (vs1,y1),SOME m1)’ \\ + rename1 ‘ltree_el (BT' X (M j1) r) q = SOME (SOME (vs1,y1),SOME m1)’ \\ Q.PAT_X_ASSUM ‘_ = SOME (vs1,y1)’ K_TAC >> gs [] \\ Q.PAT_X_ASSUM ‘_ = r1’ K_TAC \\ Q.PAT_X_ASSUM ‘_ = SOME m1’ K_TAC \\ @@ -8116,7 +8554,7 @@ Proof “y1' :string”, “Ns1 :term list”)) ‘N1’ \\ ‘TAKE (LAMl_size N0) vs1 = vs1’ by rw [] \\ POP_ASSUM (rfs o wrap) \\ - ‘subterm X (H j1) p r <> NONE’ by ASM_SIMP_TAC std_ss [] \\ + ‘subterm X (H j1) q r <> NONE’ by ASM_SIMP_TAC std_ss [] \\ Know ‘IMAGE y (count k) SUBSET X UNION RANK r1’ >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘X UNION RANK r’ \\ reverse CONJ_TAC @@ -8128,8 +8566,8 @@ Proof FIRST_X_ASSUM ACCEPT_TAC) >> DISCH_TAC \\ Know ‘set vs SUBSET X UNION RANK r1’ >- (Suff ‘set vs SUBSET RANK r1’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘RANK r’ >> art [] \\ + simp [Abbr ‘r1’, RANK_MONO]) >> DISCH_TAC \\ Know ‘set ys SUBSET X UNION RANK r1’ >- (Suff ‘set ys SUBSET RANK r1’ >- SET_TAC [] \\ qunabbrev_tac ‘ys’ \\ @@ -8142,48 +8580,55 @@ Proof Know ‘m1 <= d_max’ >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘d’ \\ reverse CONJ_TAC >- rw [Abbr ‘d_max’] \\ - Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j1) p’ \\ - reverse CONJ_TAC >- simp [] \\ + Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j1) q’ \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC subterm_width_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ qunabbrevl_tac [‘m1’, ‘N0’] \\ - ‘N = subterm' X (M j1) p r’ by rw [] >> POP_ORW \\ + ‘N = subterm' X (M j1) q r’ by rw [] >> POP_ORW \\ MATCH_MP_TAC subterm_width_last >> simp []) >> DISCH_TAC \\ - Know ‘solvable (subterm' X (H j1) p r)’ + Know ‘solvable (subterm' X (H j1) q r)’ >- (ASM_SIMP_TAC std_ss [] \\ - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0’, ‘r1’, ‘d_max’] \\ - simp [principle_hnf_tpm']) >> DISCH_TAC \\ - Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) p r <> NONE /\ _’ ASSUME_TAC \\ - MP_TAC (Q.SPECL [‘p’, ‘X’, ‘H (j1 :num)’, ‘r’] BT_subterm_thm) \\ + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r1’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + POP_ASSUM MP_TAC >> rw [Abbr ‘m1’] \\ + Know ‘y i IN X UNION RANK r’ >- METIS_TAC [SUBSET_DEF] \\ + Suff ‘RANK r SUBSET RANK r1’ >- SET_TAC [] \\ + rw [Abbr ‘r1’, RANK_MONO]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ ASSUME_TAC \\ + MP_TAC (Q.SPECL [‘q’, ‘X’, ‘H (j1 :num)’, ‘r’] BT_subterm_thm) \\ simp [] >> STRIP_TAC (* this asserts ‘x’ *) \\ NTAC 3 (Cases_on ‘x’ >> fs []) \\ - simp [head_equivalent_def, head_equivalent'_def]) + simp [head_equivalent_def]) (* stage work, now “subterm' X (M j1) p r)” and “subterm' X (M j2) p r)” are both solvable. *) - >> MP_TAC (Q.SPECL [‘p’, ‘X’, ‘M (j1 :num)’, ‘r’] BT_subterm_thm) + >> PRINT_TAC "stage work on subtree_equiv_lemma: L8620" + >> MP_TAC (Q.SPECL [‘q’, ‘X’, ‘M (j1 :num)’, ‘r’] BT_subterm_thm) >> simp [] >> STRIP_TAC (* this asserts ‘x’ *) >> NTAC 3 (Cases_on ‘x’ >> fs []) - >> qabbrev_tac ‘r1 = r + LENGTH p’ - >> rename1 ‘subterm X (M j1) p r = SOME (N,r1)’ + >> qabbrev_tac ‘r1 = r + LENGTH q’ + >> rename1 ‘subterm X (M j1) q r = SOME (N,r1)’ >> qabbrev_tac ‘N0 = principle_hnf N’ >> qabbrev_tac ‘m1 = hnf_children_size N0’ - >> rename1 ‘ltree_el (BT' X (M j1) r) p = SOME (SOME (vs1,y1),SOME m1)’ + >> rename1 ‘ltree_el (BT' X (M j1) r) q = SOME (SOME (vs1,y1),SOME m1)’ >> Q.PAT_X_ASSUM ‘_ = SOME (vs1,y1)’ K_TAC >> gs [] >> Q.PAT_X_ASSUM ‘_ = r1’ K_TAC >> Q.PAT_X_ASSUM ‘_ = SOME m1’ K_TAC >> qabbrev_tac ‘n1 = LAMl_size N0’ - >> MP_TAC (Q.SPECL [‘p’, ‘X’, ‘M (j2 :num)’, ‘r’] BT_subterm_thm) + >> MP_TAC (Q.SPECL [‘q’, ‘X’, ‘M (j2 :num)’, ‘r’] BT_subterm_thm) >> simp [] >> STRIP_TAC (* this asserts ‘x’ *) >> NTAC 3 (Cases_on ‘x’ >> fs []) - >> rename1 ‘subterm X (M j2) p r = SOME (N',r1)’ + >> rename1 ‘subterm X (M j2) q r = SOME (N',r1)’ >> qabbrev_tac ‘N0' = principle_hnf N'’ >> qabbrev_tac ‘m2 = hnf_children_size N0'’ - >> rename1 ‘ltree_el (BT' X (M j2) r) p = SOME (SOME (vs2,y2),SOME m2)’ + >> rename1 ‘ltree_el (BT' X (M j2) r) q = SOME (SOME (vs2,y2),SOME m2)’ >> Q.PAT_X_ASSUM ‘_ = SOME (vs2,y2)’ K_TAC >> gs [] >> Q.PAT_X_ASSUM ‘_ = r1’ K_TAC >> Q.PAT_X_ASSUM ‘_ = SOME m2’ K_TAC >> qabbrev_tac ‘n2 = LAMl_size N0'’ - >> simp [head_equivalent_def, head_equivalent'_def] + >> simp [head_equivalent_def] (* decompose N *) >> Q.PAT_X_ASSUM ‘RNEWS r1 n1 X = vs1’ (fs o wrap o SYM) >> Q_TAC (RNEWS_TAC (“vs1 :string list”, “r1 :num”, “n1 :num”)) ‘X’ @@ -8211,8 +8656,8 @@ Proof >> fs [] >> Q.PAT_X_ASSUM ‘y2' = y2’ (fs o wrap) >> Q.PAT_X_ASSUM ‘y1' = y1’ (fs o wrap) - >> Know ‘subterm X (H j1) p r <> NONE /\ - subterm X (H j2) p r <> NONE’ + >> Know ‘subterm X (H j1) q r <> NONE /\ + subterm X (H j2) q r <> NONE’ >- ASM_SIMP_TAC std_ss [] >> STRIP_TAC >> Know ‘IMAGE y (count k) SUBSET X UNION RANK r1’ @@ -8227,8 +8672,8 @@ Proof >> DISCH_TAC >> Know ‘set vs SUBSET X UNION RANK r1’ >- (Suff ‘set vs SUBSET RANK r1’ >- SET_TAC [] \\ - qunabbrev_tac ‘vs’ \\ - MATCH_MP_TAC RNEWS_SUBSET_RANK >> simp [Abbr ‘r1’]) + Q_TAC (TRANS_TAC SUBSET_TRANS) ‘RANK r’ >> art [] \\ + simp [Abbr ‘r1’, RANK_MONO]) >> DISCH_TAC >> Know ‘set ys SUBSET X UNION RANK r1’ >- (Suff ‘set ys SUBSET RANK r1’ >- SET_TAC [] \\ @@ -8243,9 +8688,12 @@ Proof Q.EXISTS_TAC ‘r1’ >> simp [Abbr ‘pm’, MAP_REVERSE, MAP_ZIP]) >> STRIP_TAC >> Know ‘m1 <= d’ (* m1 = hnf_children_size N0 *) - >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j1) p’ >> simp [] \\ - qunabbrevl_tac [‘m1’, ‘N0’] \\ - ‘N = subterm' X (M j1) p r’ by rw [] >> POP_ORW \\ + >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j1) q’ \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC subterm_width_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ + simp [Abbr ‘m1’, Abbr ‘N0’] \\ + ‘N = subterm' X (M j1) q r’ by rw [] >> POP_ORW \\ MATCH_MP_TAC subterm_width_last >> simp []) >> DISCH_TAC >> Know ‘m1 <= d_max’ (* m1 = hnf_children_size N0 *) @@ -8253,48 +8701,59 @@ Proof rw [Abbr ‘d_max’]) >> DISCH_TAC >> Know ‘m2 <= d’ (* m2 = hnf_children_size N0' *) - >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j2) p’ \\ - reverse CONJ_TAC >- simp [] \\ + >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_width (M j2) q’ \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC subterm_width_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ qunabbrevl_tac [‘m2’, ‘N0'’] \\ - ‘N' = subterm' X (M j2) p r’ by rw [] >> POP_ORW \\ + ‘N' = subterm' X (M j2) q r’ by rw [] >> POP_ORW \\ MATCH_MP_TAC subterm_width_last >> simp []) >> DISCH_TAC >> Know ‘m2 <= d_max’ (* m2 = hnf_children_size N0' *) >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘d’ \\ rw [Abbr ‘d_max’]) >> DISCH_TAC - >> Know ‘solvable (subterm' X (H j1) p r) /\ - solvable (subterm' X (H j2) p r)’ + >> Know ‘solvable (subterm' X (H j1) q r) /\ + solvable (subterm' X (H j2) q r)’ >- (ASM_SIMP_TAC std_ss [] \\ CONJ_TAC >| (* 2 subgoals *) [ (* goal 1 (of 2) *) - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0’, ‘r1’, ‘d_max’] \\ - simp [principle_hnf_tpm'], + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r1’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + POP_ASSUM MP_TAC >> rw [Abbr ‘m1’] \\ + Q.PAT_X_ASSUM ‘IMAGE y (count k) SUBSET X UNION RANK r1’ MP_TAC \\ + rw [SUBSET_DEF] \\ + POP_ASSUM MATCH_MP_TAC >> Q.EXISTS_TAC ‘i’ >> art [], (* goal 2 (of 2) *) - MATCH_MP_TAC solvable_ISUB_permutator \\ - qexistsl_tac [‘X’, ‘tpm (REVERSE pm) N0'’, ‘r1’, ‘d_max’] \\ - simp [principle_hnf_tpm'] ]) + MATCH_MP_TAC (cj 1 solvable_isub_permutator_alt) \\ + qexistsl_tac [‘X’, ‘r1’, ‘d_max’, ‘y’, ‘k’] \\ + simp [subterm_width_nil, principle_hnf_tpm'] \\ + POP_ASSUM MP_TAC >> rw [Abbr ‘m1’] \\ + Q.PAT_X_ASSUM ‘IMAGE y (count k) SUBSET X UNION RANK r1’ MP_TAC \\ + rw [SUBSET_DEF] \\ + POP_ASSUM MATCH_MP_TAC >> Q.EXISTS_TAC ‘i’ >> art [] ]) >> STRIP_TAC + >> PRINT_TAC "stage work on subtree_equiv_lemma: L8750" >> Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ ASSUME_TAC - >> MP_TAC (Q.SPECL [‘p’, ‘X’, ‘H (j1 :num)’, ‘r’] BT_subterm_thm) + >> MP_TAC (Q.SPECL [‘q’, ‘X’, ‘H (j1 :num)’, ‘r’] BT_subterm_thm) >> simp [] >> STRIP_TAC (* this asserts ‘x’ *) >> NTAC 3 (Cases_on ‘x’ >> fs []) - >> rename1 ‘subterm X (H j1) p r = SOME (W,r1)’ + >> rename1 ‘subterm X (H j1) q r = SOME (W,r1)’ >> qabbrev_tac ‘W0 = principle_hnf W’ >> qabbrev_tac ‘m3 = hnf_children_size W0’ - >> rename1 ‘ltree_el (BT' X (H j1) r) p = SOME (SOME (vs3,y3),SOME m3)’ + >> rename1 ‘ltree_el (BT' X (H j1) r) q = SOME (SOME (vs3,y3),SOME m3)’ >> Q.PAT_X_ASSUM ‘_ = SOME (vs3,y3)’ K_TAC >> Q.PAT_X_ASSUM ‘_ = SOME m3’ K_TAC >> qabbrev_tac ‘n3 = LAMl_size W0’ >> Q.PAT_X_ASSUM ‘_ = r1’ (fs o wrap) >> T_TAC - >> MP_TAC (Q.SPECL [‘p’, ‘X’, ‘H (j2 :num)’, ‘r’] BT_subterm_thm) + >> MP_TAC (Q.SPECL [‘q’, ‘X’, ‘H (j2 :num)’, ‘r’] BT_subterm_thm) >> simp [] >> STRIP_TAC (* this asserts ‘x’ *) >> NTAC 3 (Cases_on ‘x’ >> fs []) - >> rename1 ‘subterm X (H j2) p r = SOME (W',r1)’ + >> rename1 ‘subterm X (H j2) q r = SOME (W',r1)’ >> qabbrev_tac ‘W0' = principle_hnf W'’ >> qabbrev_tac ‘m4 = hnf_children_size W0'’ - >> rename1 ‘ltree_el (BT' X (H j2) r) p = SOME (SOME (vs4,y4),SOME m4)’ + >> rename1 ‘ltree_el (BT' X (H j2) r) q = SOME (SOME (vs4,y4),SOME m4)’ >> Q.PAT_X_ASSUM ‘_ = SOME (vs4,y4)’ K_TAC >> Q.PAT_X_ASSUM ‘_ = SOME m4’ K_TAC >> qabbrev_tac ‘n4 = LAMl_size W0'’ @@ -8324,23 +8783,25 @@ Proof >> simp [] (* y4' = y4 *) >> DISCH_THEN (rfs o wrap) (* properties of W0 *) - >> ‘LAMl_size W0 = n3 /\ hnf_children_size W0 = m3 /\ hnf_headvar W1 = y3’ - by rw [] + >> ‘LAMl_size W0 = n3 /\ hnf_children_size W0 = m3 /\ + hnf_headvar W1 = y3’ by rw [] >> Q.PAT_X_ASSUM ‘W0 = _’ (ASSUME_TAC o SYM) >> Q.PAT_X_ASSUM ‘W1 = _’ (ASSUME_TAC o SYM) (* properties of W0' *) - >> ‘LAMl_size W0' = n4 /\ hnf_children_size W0' = m4 /\ hnf_headvar W1' = y4’ - by rw [] + >> ‘LAMl_size W0' = n4 /\ hnf_children_size W0' = m4 /\ + hnf_headvar W1' = y4’ by rw [] >> Q.PAT_X_ASSUM ‘W0' = _’ (ASSUME_TAC o SYM) >> Q.PAT_X_ASSUM ‘W1' = _’ (ASSUME_TAC o SYM) - >> simp [head_equivalent_def, head_equivalent'_def] - (* stage work *) + >> simp [head_equivalent_def] + (* final shape of the goal: + y1 <> y2 \/ m2 + n1 <> m1 + n2 ==> y3 <> y4 \/ m4 + n3 <> m3 + n4 + *) >> Know ‘W = tpm (REVERSE pm) N ISUB ss’ - >- (Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) p r <> NONE /\ _’ + >- (Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ (MP_TAC o Q.SPEC ‘j1’) >> simp []) >> DISCH_TAC >> Know ‘W' = tpm (REVERSE pm) N' ISUB ss’ - >- (Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) p r <> NONE /\ _’ + >- (Q.PAT_X_ASSUM ‘!i. i < k ==> subterm X (H i) q r <> NONE /\ _’ (MP_TAC o Q.SPEC ‘j2’) >> simp []) >> DISCH_TAC (* applying hreduce_ISUB and tpm_hreduces *) @@ -8370,7 +8831,9 @@ Proof Suff ‘DISJOINT (set vs1) (set vs)’ >- (rw [DISJOINT_ALT] \\ FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) \\ - qunabbrevl_tac [‘vs1’, ‘vs’] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs1’, ‘vs0’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’], (* goal 2 (of 2) *) POP_ASSUM MP_TAC \\ @@ -8391,7 +8854,9 @@ Proof Suff ‘DISJOINT (set vs2) (set vs)’ >- (rw [DISJOINT_ALT] \\ FIRST_X_ASSUM MATCH_MP_TAC >> simp [EL_MEM]) \\ - qunabbrevl_tac [‘vs2’, ‘vs’] \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs2’, ‘vs0’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’], (* goal 2 (of 2) *) POP_ASSUM MP_TAC \\ @@ -8423,7 +8888,7 @@ Proof >> qabbrev_tac ‘Ns1'' = MAP (\t. t ISUB ss) Ns1'’ >> qabbrev_tac ‘Ns2'' = MAP (\t. t ISUB ss) Ns2'’ >> ‘LENGTH Ns1'' = m1 /\ LENGTH Ns2'' = m2’ - by simp [Abbr ‘Ns1''’, Abbr ‘Ns2''’, Abbr ‘Ns1'’, Abbr ‘Ns2'’] + by simp [Abbr ‘Ns1''’, Abbr ‘Ns2''’, Abbr ‘Ns1'’, Abbr ‘Ns2'’] (* stage work (before final case analysis) N -h->* LAMl vs1 (VAR y1 @* Ns1) (= N0) @@ -8470,7 +8935,7 @@ Proof qunabbrevl_tac [‘m3’, ‘m4’] \\ Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ - simp []) + simp [Abbr ‘y1'’, Abbr ‘y2'’]) (* Case 2 (of 4): hard, we try to directly prove ‘y1' <> y4’ N -h->* LAMl vs1 (VAR y1 @* Ns1) (= N0) @@ -8499,12 +8964,11 @@ Proof LAMl |<----------- zs2' ----------->| VAR h LAMl |<----vs2----->|<----zs2---->|h| VAR h n4 = n2 + d_max - m2 +1 - - It seems that y4 is ‘LAST vs4’ while y1' (at most in vs1/vs3, which is - strictly shorter than vs4), thus cannot be the same (ALL_DISTINCT vs4). *) - >- (POP_ASSUM MP_TAC >> simp [ISUB_VAR_FRESH'] \\ + >- (PRINT_TAC "stage work on subtree_equiv_lemma: L8981" \\ + POP_ASSUM MP_TAC >> simp [ISUB_VAR_FRESH'] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘j3’ STRIP_ASSUME_TAC) \\ + ‘(LEAST j. y j = y2') = f j3’ by rw [] >> POP_ORW \\ ONCE_REWRITE_TAC [TAUT ‘p /\ q ==> r <=> p ==> q ==> r’] >> STRIP_TAC \\ ‘hnf (LAMl vs1 (VAR y1' @* Ns1''))’ by rw [hnf_appstar] \\ ‘LAMl vs1 (VAR y1' @* Ns1'') = W0’ by METIS_TAC [principle_hnf_thm'] \\ @@ -8521,9 +8985,9 @@ Proof qabbrev_tac ‘X' = set vs4 UNION FV W1' UNION BIGUNION (IMAGE FV (set Ns2''))’ \\ ‘FINITE X'’ by rw [Abbr ‘X'’] \\ - qabbrev_tac ‘d' = MAX n4 (SUC d_max)’ \\ + qabbrev_tac ‘d' = MAX n4 (SUC (d_max' j3))’ \\ Q_TAC (NEWS_TAC (“L :string list”, “d' :num”)) ‘X'’ \\ - ‘SUC d_max <= LENGTH L /\ n4 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ + ‘d_max' j3 < LENGTH L /\ n4 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ Know ‘DISJOINT (set L) (set vs2) /\ DISJOINT (set L) (set vs4)’ >- (rw [Abbr ‘L’, Abbr ‘vs2’, Abbr ‘vs4’] (* 2 subgoals, same tactics *) \\ @@ -8533,11 +8997,13 @@ Proof qunabbrev_tac ‘X'’ \\ DISCH_THEN (STRIP_ASSUME_TAC o REWRITE_RULE [DISJOINT_UNION']) \\ STRIP_TAC (* W -h->* ... *) \\ + ‘m2 <= d_max' j3’ by simp [Abbr ‘d_max'’] \\ (* applying hreduce_permutator_shared *) - MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max’, ‘L’] hreduce_permutator_shared) \\ - simp [] \\ + MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max + f (j3 :num)’, ‘L’] + hreduce_permutator_shared) >> simp [] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘zs2’ (Q.X_CHOOSE_THEN ‘z2’ STRIP_ASSUME_TAC)) \\ - Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ MP_TAC \\ + qabbrev_tac ‘P' = P (f j3)’ \\ + Q.PAT_X_ASSUM ‘P' @* Ns2'' -h->* _’ MP_TAC \\ qabbrev_tac ‘h = LAST L’ (* the new shared head variable *) \\ qabbrev_tac ‘L' = FRONT L’ \\ ‘L <> []’ by rw [GSYM LENGTH_NON_NIL] \\ @@ -8554,10 +9020,10 @@ Proof POP_ORW \\ simp [IS_SUFFIX, Abbr ‘xs2’]) >> DISCH_TAC \\ ‘ALL_DISTINCT xs2’ by PROVE_TAC [IS_SUFFIX_ALL_DISTINCT] \\ - Know ‘LAMl vs2 (P @* Ns2'') -h->* + Know ‘LAMl vs2 (P' @* Ns2'') -h->* LAMl vs2 (LAMl zs2 (LAM h (VAR h @* Ns2'' @* MAP VAR zs2)))’ >- simp [hreduce_LAMl] \\ - Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ K_TAC \\ + Q.PAT_X_ASSUM ‘P' @* Ns2'' -h->* _’ K_TAC \\ REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND] \\ qabbrev_tac ‘Ns2x = Ns2'' ++ MAP VAR zs2’ \\ REWRITE_TAC [GSYM LAMl_SNOC] \\ @@ -8573,7 +9039,7 @@ Proof >- (Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ simp []) >> DISCH_TAC \\ - Know ‘SUC d_max + n2 - m2 = n4’ + Know ‘SUC (d_max' j3) + n2 - m2 = n4’ >- (POP_ASSUM (REWRITE_TAC o wrap o SYM) \\ simp [Abbr ‘zs2'’]) >> DISCH_TAC \\ Know ‘vs2 <<= vs4’ @@ -8604,9 +9070,9 @@ Proof |<----------- zs2' ----------->| LAMl |<----vs2----->|<--- zs2 --->|h| VAR h @* Ns2x = Ns2'' ++ zs2 |<---- xs2 ---->| - n4 = n2 + d_max - m2 + 1 + n4 = n2 + d_max' - m2 + 1 m4 = LENGTH Ns2x = LENGTH Ns2'' + LENGTH zs2 - = m2 + d_max - m2 = d_max + = m2 + d_max - m2 = d_max' (m1 + n1 =) d_max + n1 = m1 + n2 + d_max - m2 + 1 (= m3 + n4) n1 = m1 + n2 - m2 + 1 @@ -8632,14 +9098,16 @@ Proof Know ‘DISJOINT (set vs) (set ys2)’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘vs’, ‘vs4’] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘vs4’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ Know ‘DISJOINT (set ys) (set ys2)’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ qunabbrevl_tac [‘ys’, ‘vs4’] \\ MATCH_MP_TAC DISJOINT_RNEWS \\ - ‘0 < LENGTH p’ by rw [LENGTH_NON_NIL] \\ + ‘0 < LENGTH q’ by rw [LENGTH_NON_NIL] \\ simp [Abbr ‘r1’]) >> DISCH_TAC \\ Know ‘LAMl xs2 t = LAMl ys2 (pm2 ' t)’ >- (simp [Abbr ‘pm2’, fromPairs_def] \\ @@ -8733,20 +9201,19 @@ Proof cannot be the same with y3 since ‘ALL_DISTINCT vs4’. *) Know ‘n1 <= n_max’ (* n1 = LAMl_size N0 *) - >- (qunabbrev_tac ‘n_max’ \\ - Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_length (M j1) p’ \\ + >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_length (M j1) q’ \\ reverse CONJ_TAC - >- (MATCH_MP_TAC MAX_LIST_PROPERTY >> simp [MEM_MAP] \\ - Q.EXISTS_TAC ‘M j1’ >> art [] \\ - simp [Abbr ‘M’, EL_MEM]) \\ + >- (MATCH_MP_TAC subterm_length_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ qunabbrevl_tac [‘n1’, ‘N0’] \\ - ‘N = subterm' X (M j1) p r’ by rw [] >> POP_ORW \\ + ‘N = subterm' X (M j1) q r’ by rw [] >> POP_ORW \\ MATCH_MP_TAC subterm_length_last >> simp []) >> DISCH_TAC \\ Know ‘n1 < n4’ >- (Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘n_max’ >> art [] \\ - Q.PAT_X_ASSUM ‘SUC d_max + n2 - m2 = n4’ (REWRITE_TAC o wrap o SYM) \\ - Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘d_max + n2 - m2’ \\ - simp [Abbr ‘d_max’]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘SUC (d_max' j3) + n2 - m2 = n4’ + (REWRITE_TAC o wrap o SYM) \\ + Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘d_max' j3 + n2 - m2’ \\ + simp [Abbr ‘d_max’, Abbr ‘d_max'’]) >> DISCH_TAC \\ (* applying subterm_headvar_lemma' on W *) Know ‘hnf_headvar W1 IN FV W UNION set vs3’ >- (MATCH_MP_TAC subterm_headvar_lemma' \\ @@ -8803,8 +9270,10 @@ Proof simp [ALL_DISTINCT_APPEND] \\ DISJ2_TAC >> Q.EXISTS_TAC ‘y1'’ >> art [] ]) (* Case 3 (of 4): (almost) symmetric with previous Case 2 *) - >- (Q.PAT_X_ASSUM ‘~(y1' NOTIN DOM ss)’ MP_TAC >> simp [ISUB_VAR_FRESH'] \\ + >- (PRINT_TAC "stage work on subtree_equiv_lemma: L9286" \\ + Q.PAT_X_ASSUM ‘~(y1' NOTIN DOM ss)’ MP_TAC >> simp [ISUB_VAR_FRESH'] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘j3’ STRIP_ASSUME_TAC) \\ + ‘(LEAST j. y j = y1') = f j3’ by rw [] >> POP_ORW \\ ONCE_REWRITE_TAC [TAUT ‘p /\ q ==> r <=> q ==> p ==> r’] >> STRIP_TAC \\ ‘hnf (LAMl vs2 (VAR y2' @* Ns2''))’ by rw [hnf_appstar] \\ ‘LAMl vs2 (VAR y2' @* Ns2'') = W0'’ by METIS_TAC [principle_hnf_thm'] \\ @@ -8819,9 +9288,9 @@ Proof qabbrev_tac ‘X' = set vs3 UNION FV W1 UNION BIGUNION (IMAGE FV (set Ns1''))’ \\ ‘FINITE X'’ by rw [Abbr ‘X'’] \\ - qabbrev_tac ‘d' = MAX n3 (SUC d_max)’ \\ + qabbrev_tac ‘d' = MAX n3 (SUC (d_max' j3))’ \\ Q_TAC (NEWS_TAC (“L :string list”, “d' :num”)) ‘X'’ \\ - ‘SUC d_max <= LENGTH L /\ n3 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ + ‘d_max' j3 < LENGTH L /\ n3 <= LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] \\ Know ‘DISJOINT (set L) (set vs1) /\ DISJOINT (set L) (set vs3)’ >- (rw [Abbr ‘L’, Abbr ‘vs1’, Abbr ‘vs3’] (* 2 subgoals, same tactics *) \\ @@ -8831,10 +9300,12 @@ Proof qunabbrev_tac ‘X'’ \\ DISCH_THEN (STRIP_ASSUME_TAC o REWRITE_RULE [DISJOINT_UNION']) \\ STRIP_TAC (* W -h->* ... *) \\ - MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max’, ‘L’] hreduce_permutator_shared) \\ - simp [] \\ + ‘m1 <= d_max' j3’ by simp [Abbr ‘d_max'’] \\ + MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max + f (j3 :num)’, ‘L’] + hreduce_permutator_shared) >> simp [] \\ DISCH_THEN (Q.X_CHOOSE_THEN ‘zs1’ (Q.X_CHOOSE_THEN ‘z1’ STRIP_ASSUME_TAC)) \\ - Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ MP_TAC \\ + qabbrev_tac ‘P' = P (f j3)’ \\ + Q.PAT_X_ASSUM ‘P' @* Ns1'' -h->* _’ MP_TAC \\ qabbrev_tac ‘h = LAST L’ (* the new shared head variable *) \\ qabbrev_tac ‘L' = FRONT L’ \\ ‘L <> []’ by rw [GSYM LENGTH_NON_NIL] \\ @@ -8850,10 +9321,10 @@ Proof by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] \\ POP_ORW >> simp [IS_SUFFIX, Abbr ‘xs1’]) >> DISCH_TAC \\ ‘ALL_DISTINCT xs1’ by PROVE_TAC [IS_SUFFIX_ALL_DISTINCT] \\ - Know ‘LAMl vs1 (P @* Ns1'') -h->* + Know ‘LAMl vs1 (P' @* Ns1'') -h->* LAMl vs1 (LAMl zs1 (LAM h (VAR h @* Ns1'' @* MAP VAR zs1)))’ >- simp [hreduce_LAMl] \\ - Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ K_TAC \\ + Q.PAT_X_ASSUM ‘P' @* Ns1'' -h->* _’ K_TAC \\ REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND] \\ qabbrev_tac ‘Ns1x = Ns1'' ++ MAP VAR zs1’ \\ REWRITE_TAC [GSYM LAMl_SNOC] \\ @@ -8870,7 +9341,7 @@ Proof simp []) >> DISCH_TAC \\ (* abandon ‘y1 <> y2 \/ m2 + n1 <> m1 + n2’ *) DISCH_THEN K_TAC >> DISJ1_TAC \\ - Know ‘SUC d_max + n1 - m1 = n3’ + Know ‘SUC (d_max' j3) + n1 - m1 = n3’ >- (POP_ASSUM (REWRITE_TAC o wrap o SYM) \\ simp [Abbr ‘zs1'’]) >> DISCH_TAC \\ Know ‘vs1 <<= vs3’ @@ -8882,20 +9353,315 @@ Proof Q.PAT_X_ASSUM ‘n1' <= n3’ MP_TAC \\ Q.PAT_X_ASSUM ‘vs1 = TAKE n1' vs3’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ POP_ORW >> NTAC 2 DISCH_TAC \\ - ‘zs1' = vs1 ++ xs1’ by simp [Abbr ‘zs1'’, Abbr ‘xs1’] \\ + ‘zs1' = vs1 ++ xs1’ by simp [Abbr ‘zs1'’, Abbr ‘xs1’] \\ + qabbrev_tac ‘ys1 = DROP n1 vs3’ \\ + ‘ALL_DISTINCT ys1’ by simp [Abbr ‘ys1’, ALL_DISTINCT_DROP] \\ + (* Structure of W0: + + LAMl |<---(vs1)--- vs3 ---(ys1)---->| VAR y3 (= LAST vs3) in ROW r1 + -------------------------------------------------------------- + |<----------- zs1' ----------->| + LAMl |<----vs1----->|<--- zs1 --->|h| VAR h @* Ns1x = Ns1'' ++ zs1 + |<---- xs1 ---->| + + LAMl vs3 (VAR y3 @* Ns3) = W0 + LAMl zs1' (VAR h @* Ns1x) = W0 + *) + Know ‘DISJOINT (set xs1) (set ys1)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs3’ \\ + reverse CONJ_TAC >- simp [LIST_TO_SET_DROP, Abbr ‘ys1’] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set L’ >> art [] \\ + MATCH_MP_TAC LIST_TO_SET_SUFFIX >> art []) >> DISCH_TAC \\ + ‘LENGTH xs1 = LENGTH ys1’ by simp [Abbr ‘xs1’, Abbr ‘ys1’] \\ + Know ‘LAMl vs3 (VAR y3 @* Ns3) = LAMl zs1' (VAR h @* Ns1x)’ >- rw [] \\ + ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + Q.PAT_X_ASSUM ‘zs1' = _’ (ONCE_REWRITE_TAC o wrap) \\ + simp [LAMl_APPEND] \\ + qabbrev_tac ‘t = VAR h @* Ns1x’ \\ + (* applying LAMl_ALPHA_ssub *) + qabbrev_tac ‘pm1 = fromPairs xs1 (MAP VAR ys1)’ \\ + (* NOTE: The following disjointness hold for names from different rows *) + Know ‘DISJOINT (set vs) (set ys1)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘vs3’] \\ + MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Know ‘DISJOINT (set ys) (set ys1)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ + qunabbrevl_tac [‘ys’, ‘vs3’] \\ + MATCH_MP_TAC DISJOINT_RNEWS \\ + ‘0 < LENGTH q’ by rw [LENGTH_NON_NIL] \\ + simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Know ‘LAMl xs1 t = LAMl ys1 (pm1 ' t)’ + >- (simp [Abbr ‘pm1’, fromPairs_def] \\ + MATCH_MP_TAC LAMl_ALPHA_ssub >> art [] \\ + simp [DISJOINT_UNION'] \\ + CONJ_TAC >- rw [Once DISJOINT_SYM] \\ + simp [Abbr ‘t’, Abbr ‘Ns1x’, appstar_APPEND, FV_appstar_MAP_VAR] \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set xs1’ >> art [] \\ + simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] >> SET_TAC []) \\ + simp [FV_appstar] \\ + CONJ_TAC + >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs1) (set ys1)’ MP_TAC \\ + rw [Abbr ‘xs1’, LIST_TO_SET_SNOC, DISJOINT_ALT]) \\ + simp [MEM_EL] >> rpt STRIP_TAC \\ + Q.PAT_X_ASSUM ‘_ = FV x’ (REWRITE_TAC o wrap) >> POP_ORW \\ + rename1 ‘i < m1’ >> POP_ASSUM MP_TAC \\ + simp [Abbr ‘Ns1''’, EL_MAP] >> DISCH_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘FV (EL i Ns1')’ \\ + reverse CONJ_TAC + >- (MP_TAC (Q.SPECL [‘ss’, ‘EL i Ns1'’] FV_ISUB_upperbound) \\ + simp [EL_MAP, Abbr ‘Ns1'’]) \\ + simp [Abbr ‘Ns1'’, EL_listpm, Abbr ‘pm’, REVERSE_ZIP] \\ + (* applying FV_tpm_disjoint *) + ONCE_REWRITE_TAC [DISJOINT_SYM] \\ + MATCH_MP_TAC FV_tpm_disjoint \\ + simp [ALL_DISTINCT_REVERSE] \\ + (* goal: DISJOINT (set ys1) (FV (EL i Ns1)) *) + Know ‘FV N0 SUBSET X UNION RANK r1’ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV N’ >> art [] \\ + qunabbrev_tac ‘N0’ \\ + MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) >> DISCH_TAC \\ + (* applying FV_subterm_lemma *) + Know ‘FV (EL i Ns1) SUBSET FV N UNION set vs1’ + >- (MATCH_MP_TAC FV_subterm_lemma \\ + qexistsl_tac [‘X’, ‘r1’, ‘N0’, ‘n1’, ‘m1’, ‘N1’] >> simp []) \\ + DISCH_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘FV N UNION set vs1’ >> art [] \\ + REWRITE_TAC [DISJOINT_UNION'] \\ + reverse CONJ_TAC + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ + ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘X UNION RANK r1’ >> art [] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs3’ \\ + reverse CONJ_TAC + >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + simp [SUBSET_DEF]) \\ + simp [DISJOINT_UNION'] \\ + qunabbrev_tac ‘vs3’ \\ + MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> Rewr' \\ + ‘FDOM pm1 = set xs1’ by simp [Abbr ‘pm1’, FDOM_fromPairs] \\ + ‘MEM h xs1’ by simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] \\ + simp [Abbr ‘t’, ssub_appstar] \\ + Know ‘pm1 ' h = VAR (LAST ys1)’ + >- (‘h = LAST xs1’ by rw [Abbr ‘xs1’, LAST_SNOC] >> POP_ORW \\ + ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ + ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + simp [Abbr ‘pm1’, LAST_EL] \\ + qabbrev_tac ‘j4 = PRE (LENGTH ys1)’ \\ + ‘0 < LENGTH ys1’ by rw [LENGTH_NON_NIL] \\ + ‘j4 < LENGTH ys1’ by rw [Abbr ‘j4’] \\ + ‘VAR (EL j4 ys1) = EL j4 (MAP VAR ys1)’ by simp [EL_MAP] >> POP_ORW \\ + MATCH_MP_TAC fromPairs_FAPPLY_EL >> simp []) >> Rewr' \\ + Q.PAT_X_ASSUM ‘_ = W1’ (ONCE_REWRITE_TAC o wrap o SYM) \\ + simp [] \\ + Know ‘LAST ys1 = LAST vs3’ + >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ + ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + rw [LAST_APPEND_NOT_NIL]) >> Rewr' \\ + STRIP_TAC (* y4 = LAST vs4, etc. *) \\ + (* stage work *) + Know ‘n2 <= n_max’ (* n2 = LAMl_size N0' *) + >- (Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_length (M j2) q’ \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC subterm_length_inclusive \\ + Q.EXISTS_TAC ‘p’ >> simp []) \\ + qunabbrevl_tac [‘n2’, ‘N0'’] \\ + ‘N' = subterm' X (M j2) q r’ by rw [] >> POP_ORW \\ + MATCH_MP_TAC subterm_length_last >> simp []) >> DISCH_TAC \\ + Know ‘n2 < n3’ + >- (Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘n_max’ >> art [] \\ + Q.PAT_X_ASSUM ‘SUC (d_max' j3) + n1 - m1 = n3’ + (REWRITE_TAC o wrap o SYM) \\ + Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘d_max' j3 + n1 - m1’ \\ + simp [Abbr ‘d_max’, Abbr ‘d_max'’]) >> DISCH_TAC \\ + (* applying subterm_headvar_lemma' on W *) + Know ‘hnf_headvar W1' IN FV W' UNION set vs4’ + >- (MATCH_MP_TAC subterm_headvar_lemma' \\ + qexistsl_tac [‘X’, ‘r1’, ‘W0'’, ‘n4’] >> simp []) \\ + Know ‘hnf_headvar W1' = y4’ + >- (Q.PAT_X_ASSUM ‘_ = W1'’ (REWRITE_TAC o wrap o SYM) \\ + simp []) >> Rewr' >> art [] (* y3 -> y1' *) \\ + DISCH_TAC \\ + (* NOTE: FV W' SUBSET X UNION RANK r1 *) + qabbrev_tac ‘X' = X UNION RANK r1’ \\ + Know ‘y2' IN X' UNION set vs4’ + >- (Q.PAT_X_ASSUM ‘y2' IN _’ MP_TAC \\ + Q.PAT_X_ASSUM ‘FV W' SUBSET _’ MP_TAC \\ + SET_TAC []) \\ + Q.PAT_X_ASSUM ‘y2' IN _’ K_TAC >> DISCH_TAC \\ + qunabbrev_tac ‘X'’ \\ + ‘0 < n3 /\ vs3 <> []’ by simp [GSYM LENGTH_NON_NIL] \\ + simp [LAST_EL] \\ + Q.PAT_X_ASSUM ‘y2' IN _’ MP_TAC >> simp [IN_UNION] \\ + STRIP_TAC >| (* 3 subgoals *) + [ (* goal 1 (of 3) *) + SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ + Know ‘MEM y2' vs3’ + >- (POP_ORW >> simp [EL_MEM]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘DISJOINT (set vs3) X’ MP_TAC \\ + simp [DISJOINT_ALT'] >> Q.EXISTS_TAC ‘y2'’ >> art [], + (* goal 2 (of 3) *) + SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ + Know ‘MEM y2' vs3’ + >- (POP_ORW >> simp [EL_MEM]) >> DISCH_TAC \\ + Know ‘DISJOINT (set vs3) (RANK r1)’ + >- simp [Abbr ‘vs3’, DISJOINT_RNEWS_RANK'] \\ + simp [DISJOINT_ALT] \\ + Q.EXISTS_TAC ‘y2'’ >> art [], + (* goal 3 (of 3) *) + Know ‘vs4 <<= vs3’ + >- (qunabbrevl_tac [‘vs3’, ‘vs4’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_APPEND] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘ls’ STRIP_ASSUME_TAC) \\ + Know ‘LENGTH ls = n3 - n2’ + >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ + simp []) >> DISCH_TAC \\ + ‘ls <> []’ by simp [GSYM LENGTH_NON_NIL] \\ + Know ‘EL (PRE n3) vs3 = LAST ls’ + >- (Q.PAT_X_ASSUM ‘vs3 = vs4 ++ ls’ (REWRITE_TAC o wrap) \\ + simp [EL_APPEND2, LAST_EL, PRE_SUB]) >> Rewr' \\ + SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ + Know ‘MEM y2' ls’ + >- (POP_ORW >> simp [MEM_LAST_NOT_NIL]) >> DISCH_TAC \\ + Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs3 = vs4 ++ ls’ (REWRITE_TAC o wrap) \\ + simp [ALL_DISTINCT_APPEND] \\ + DISJ2_TAC >> Q.EXISTS_TAC ‘y2'’ >> art [] ]) + (* Case 4 (of 4): hardest *) + >> PRINT_TAC "stage work on subtree_equiv_lemma: L9559" + >> NTAC 2 (POP_ASSUM (MP_TAC o REWRITE_RULE [])) + >> simp [] + >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j3’ STRIP_ASSUME_TAC) + >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j4’ STRIP_ASSUME_TAC) + >> qabbrev_tac ‘X' = set vs3 UNION set vs4 UNION + FV W1 UNION FV W1' UNION + BIGUNION (IMAGE FV (set Ns1'')) UNION + BIGUNION (IMAGE FV (set Ns2''))’ + >> ‘FINITE X'’ by rw [Abbr ‘X'’] + >> qabbrev_tac ‘d' = MAX (MAX n3 n4) + (MAX (SUC (d_max' j3)) + (SUC (d_max' j4)))’ + >> Q_TAC (NEWS_TAC (“L :string list”, “d' :num”)) ‘X'’ + >> ‘n3 <= LENGTH L /\ n4 <= LENGTH L /\ + d_max' j3 < LENGTH L /\ d_max' j4 < LENGTH L’ by simp [Abbr ‘d'’, MAX_LE] + >> Know ‘DISJOINT (set L) (set vs1) /\ + DISJOINT (set L) (set vs2) /\ + DISJOINT (set L) (set vs3) /\ + DISJOINT (set L) (set vs4)’ + >- (rw [Abbr ‘L’, Abbr ‘vs1’, Abbr ‘vs2’, Abbr ‘vs3’, Abbr ‘vs4’] \\ + MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) + >> STRIP_TAC + >> Q.PAT_X_ASSUM ‘FINITE X'’ K_TAC + >> Q.PAT_X_ASSUM ‘DISJOINT (set L) X'’ MP_TAC + >> qunabbrev_tac ‘X'’ + >> DISCH_THEN (STRIP_ASSUME_TAC o REWRITE_RULE [DISJOINT_UNION']) + >> simp [] (* VAR (y j3) ISUB ss = P (f j3), etc. *) + >> STRIP_TAC (* W -h->* ... *) + >> ‘m1 <= d_max' j3 /\ m2 <= d_max' j4’ by simp [Abbr ‘d_max'’] + (* applying hreduce_permutator_shared *) + >> MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max + f (j3 :num)’, ‘L’] hreduce_permutator_shared) + >> simp [] + >> DISCH_THEN (Q.X_CHOOSE_THEN ‘zs1’ (Q.X_CHOOSE_THEN ‘z1’ STRIP_ASSUME_TAC)) + (* applying hreduce_permutator_shared again *) + >> MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max + f (j4 :num)’, ‘L’] hreduce_permutator_shared) + >> simp [] + >> DISCH_THEN (Q.X_CHOOSE_THEN ‘zs2’ (Q.X_CHOOSE_THEN ‘z2’ STRIP_ASSUME_TAC)) + >> qabbrev_tac ‘P3 = P (f j3)’ + >> qabbrev_tac ‘P4 = P (f j4)’ + >> Q.PAT_X_ASSUM ‘P3 @* Ns1'' -h->* _’ MP_TAC + >> Q.PAT_X_ASSUM ‘P4 @* Ns2'' -h->* _’ MP_TAC + >> qabbrev_tac ‘h = LAST L’ (* The new shared head variable *) + >> qabbrev_tac ‘L' = FRONT L’ + >> ‘L <> []’ by rw [GSYM LENGTH_NON_NIL] + >> NTAC 2 (Q.PAT_X_ASSUM ‘IS_SUFFIX L _’ MP_TAC) + >> ‘L = SNOC h L'’ by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] + >> POP_ORW >> simp [IS_SUFFIX] + >> NTAC 2 STRIP_TAC + >> Q.PAT_X_ASSUM ‘z1 = z2’ (simp o wrap o SYM) + >> Q.PAT_X_ASSUM ‘h = z1’ (simp o wrap o SYM) + >> NTAC 2 DISCH_TAC + >> qabbrev_tac ‘xs1 = SNOC h zs1’ (* suffix of L *) + >> qabbrev_tac ‘xs2 = SNOC h zs2’ (* suffix of L *) + >> Know ‘IS_SUFFIX L xs1 /\ IS_SUFFIX L xs2’ + >- (‘L = SNOC h L'’ + by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] \\ + POP_ORW >> simp [IS_SUFFIX, Abbr ‘xs1’, Abbr ‘xs2’]) + >> STRIP_TAC + >> ‘ALL_DISTINCT xs1 /\ ALL_DISTINCT xs2’ by PROVE_TAC [IS_SUFFIX_ALL_DISTINCT] + >> Know ‘LAMl vs1 (P3 @* Ns1'') -h->* + LAMl vs1 (LAMl zs1 (LAM h (VAR h @* Ns1'' @* MAP VAR zs1))) /\ + LAMl vs2 (P4 @* Ns2'') -h->* + LAMl vs2 (LAMl zs2 (LAM h (VAR h @* Ns2'' @* MAP VAR zs2)))’ + >- simp [hreduce_LAMl] + >> Q.PAT_X_ASSUM ‘P3 @* Ns1'' -h->* _’ K_TAC + >> Q.PAT_X_ASSUM ‘P4 @* Ns2'' -h->* _’ K_TAC + >> REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND] + >> qabbrev_tac ‘Ns1x = Ns1'' ++ MAP VAR zs1’ + >> qabbrev_tac ‘Ns2x = Ns2'' ++ MAP VAR zs2’ + >> REWRITE_TAC [GSYM LAMl_SNOC] + >> qabbrev_tac ‘zs1' = SNOC h (vs1 ++ zs1)’ + >> qabbrev_tac ‘zs2' = SNOC h (vs2 ++ zs2)’ + >> STRIP_TAC + >> Know ‘W -h->* LAMl zs1' (VAR h @* Ns1x) /\ + W' -h->* LAMl zs2' (VAR h @* Ns2x)’ + >- PROVE_TAC [hreduce_TRANS] + >> NTAC 2 (POP_ASSUM K_TAC) >> STRIP_TAC + >> qunabbrevl_tac [‘P3’, ‘P4’] + >> Know ‘LAMl zs1' (VAR h @* Ns1x) = W0 /\ + LAMl zs2' (VAR h @* Ns2x) = W0'’ + >- (‘hnf (LAMl zs1' (VAR h @* Ns1x)) /\ hnf (LAMl zs2' (VAR h @* Ns2x))’ + by rw [hnf_appstar] \\ + METIS_TAC [principle_hnf_thm']) + >> STRIP_TAC + >> Know ‘LENGTH zs1' = n3 /\ LENGTH zs2' = n4’ + >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) >> simp []) + >> STRIP_TAC + (* key equations about n3 and n4, m3 and m4 *) + >> Know ‘d_max' j3 = m3 /\ d_max' j4 = m4’ + >- (Q.PAT_X_ASSUM ‘_ = m3’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = m4’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) \\ + simp [Abbr ‘Ns1x’, Abbr ‘Ns2x’]) + >> STRIP_TAC + >> Know ‘SUC (d_max' j3) + n1 - m1 = n3 /\ + SUC (d_max' j4) + n2 - m2 = n4’ + >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ + simp [Abbr ‘zs1'’, Abbr ‘zs2'’]) + >> simp [] >> STRIP_TAC + (* stage work *) + >> Know ‘LAST vs3 = y3’ + >- (Know ‘vs1 <<= vs3’ + >- (qunabbrevl_tac [‘vs1’, ‘vs3’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n1'’ STRIP_ASSUME_TAC) \\ + ‘n1' = n1’ by rw [] \\ + Q.PAT_X_ASSUM ‘n1' <= n3’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs1 = TAKE n1' vs3’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> NTAC 2 DISCH_TAC \\ + ‘vs1 ++ xs1 = zs1'’ by simp [Abbr ‘zs1'’, Abbr ‘xs1’] \\ qabbrev_tac ‘ys1 = DROP n1 vs3’ \\ ‘ALL_DISTINCT ys1’ by simp [Abbr ‘ys1’, ALL_DISTINCT_DROP] \\ - (* Structure of W0: - - LAMl |<---(vs1)--- vs3 ---(ys1)---->| VAR y3 (= LAST vs3) in ROW r1 - -------------------------------------------------------------- - |<----------- zs1' ----------->| - LAMl |<----vs1----->|<--- zs1 --->|h| VAR h @* Ns1x = Ns1'' ++ zs1 - |<---- xs1 ---->| - - LAMl vs3 (VAR y3 @* Ns3) = W0 - LAMl zs1' (VAR h @* Ns1x) = W0 - *) Know ‘DISJOINT (set xs1) (set ys1)’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs3’ \\ @@ -8906,7 +9672,7 @@ Proof ‘LENGTH xs1 = LENGTH ys1’ by simp [Abbr ‘xs1’, Abbr ‘ys1’] \\ Know ‘LAMl vs3 (VAR y3 @* Ns3) = LAMl zs1' (VAR h @* Ns1x)’ >- rw [] \\ ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - Q.PAT_X_ASSUM ‘zs1' = _’ (ONCE_REWRITE_TAC o wrap) \\ + Q.PAT_X_ASSUM ‘_ = zs1'’ (ONCE_REWRITE_TAC o wrap o SYM) \\ simp [LAMl_APPEND] \\ qabbrev_tac ‘t = VAR h @* Ns1x’ \\ (* applying LAMl_ALPHA_ssub *) @@ -8915,14 +9681,16 @@ Proof Know ‘DISJOINT (set vs) (set ys1)’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘vs’, ‘vs3’] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘vs3’] \\ MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ Know ‘DISJOINT (set ys) (set ys1)’ >- (MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ qunabbrevl_tac [‘ys’, ‘vs3’] \\ MATCH_MP_TAC DISJOINT_RNEWS \\ - ‘0 < LENGTH p’ by rw [LENGTH_NON_NIL] \\ + ‘0 < LENGTH q’ by rw [LENGTH_NON_NIL] \\ simp [Abbr ‘r1’]) >> DISCH_TAC \\ Know ‘LAMl xs1 t = LAMl ys1 (pm1 ' t)’ >- (simp [Abbr ‘pm1’, fromPairs_def] \\ @@ -8966,248 +9734,166 @@ Proof Q.EXISTS_TAC ‘FV N UNION set vs1’ >> art [] \\ REWRITE_TAC [DISJOINT_UNION'] \\ reverse CONJ_TAC - >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ - ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ + ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘X UNION RANK r1’ >> art [] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs3’ \\ + reverse CONJ_TAC + >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + simp [SUBSET_DEF]) \\ + simp [DISJOINT_UNION'] \\ + qunabbrev_tac ‘vs3’ \\ + MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> Rewr' \\ + ‘vs1 ++ ys1 = vs3’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + ‘FDOM pm1 = set xs1’ by simp [Abbr ‘pm1’, FDOM_fromPairs] \\ + ‘MEM h xs1’ by simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] \\ + simp [Abbr ‘t’, ssub_appstar] \\ + Know ‘pm1 ' h = VAR (LAST ys1)’ + >- (‘h = LAST xs1’ by rw [Abbr ‘xs1’, LAST_SNOC] >> POP_ORW \\ + ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ + ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + simp [Abbr ‘pm1’, LAST_EL] \\ + qabbrev_tac ‘j5 = PRE (LENGTH ys1)’ \\ + ‘0 < LENGTH ys1’ by rw [LENGTH_NON_NIL] \\ + ‘j5 < LENGTH ys1’ by rw [Abbr ‘j5’] \\ + ‘VAR (EL j5 ys1) = EL j5 (MAP VAR ys1)’ by simp [EL_MAP] >> POP_ORW \\ + MATCH_MP_TAC fromPairs_FAPPLY_EL >> simp []) >> Rewr' \\ + Q.PAT_X_ASSUM ‘_ = W1’ (ONCE_REWRITE_TAC o wrap o SYM) \\ + Know ‘LAST ys1 = LAST vs3’ + >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ + ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + rw [LAST_APPEND_NOT_NIL]) >> Rewr' \\ + simp []) + >> DISCH_TAC + >> Know ‘LAST vs4 = y4’ + >- (Know ‘vs2 <<= vs4’ + >- (qunabbrevl_tac [‘vs2’, ‘vs4’] \\ + MATCH_MP_TAC RNEWS_prefix >> simp []) \\ + simp [IS_PREFIX_EQ_TAKE] \\ + DISCH_THEN (Q.X_CHOOSE_THEN ‘n2'’ STRIP_ASSUME_TAC) \\ + ‘n2' = n2’ by rw [] \\ + Q.PAT_X_ASSUM ‘n2' <= n4’ MP_TAC \\ + Q.PAT_X_ASSUM ‘vs2 = TAKE n2' vs4’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ + POP_ORW >> NTAC 2 DISCH_TAC \\ + ‘vs2 ++ xs2 = zs2'’ by simp [Abbr ‘zs2'’, Abbr ‘xs2’] \\ + qabbrev_tac ‘ys2 = DROP n2 vs4’ \\ + ‘ALL_DISTINCT ys2’ by simp [Abbr ‘ys2’, ALL_DISTINCT_DROP] \\ + Know ‘DISJOINT (set xs2) (set ys2)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs4’ \\ + reverse CONJ_TAC >- simp [LIST_TO_SET_DROP, Abbr ‘ys2’] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set L’ >> art [] \\ + MATCH_MP_TAC LIST_TO_SET_SUFFIX >> art []) >> DISCH_TAC \\ + ‘LENGTH xs2 = LENGTH ys2’ by simp [Abbr ‘xs2’, Abbr ‘ys2’] \\ + Know ‘LAMl vs4 (VAR y4 @* Ns4) = LAMl zs2' (VAR h @* Ns2x)’ >- rw [] \\ + ‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + Q.PAT_X_ASSUM ‘_ = zs2'’ (ONCE_REWRITE_TAC o wrap o SYM) \\ + simp [LAMl_APPEND] \\ + qabbrev_tac ‘t = VAR h @* Ns2x’ \\ + (* applying LAMl_ALPHA_ssub *) + qabbrev_tac ‘pm2 = fromPairs xs2 (MAP VAR ys2)’ \\ + Know ‘DISJOINT (set vs) (set ys2)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set vs0’ >> art [] \\ + qunabbrevl_tac [‘vs0’, ‘vs4’] \\ + MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Know ‘DISJOINT (set ys) (set ys2)’ + >- (MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ + qunabbrevl_tac [‘ys’, ‘vs4’] \\ + MATCH_MP_TAC DISJOINT_RNEWS \\ + ‘0 < LENGTH q’ by rw [LENGTH_NON_NIL] \\ + simp [Abbr ‘r1’]) >> DISCH_TAC \\ + Know ‘LAMl xs2 t = LAMl ys2 (pm2 ' t)’ + >- (simp [Abbr ‘pm2’, fromPairs_def] \\ + MATCH_MP_TAC LAMl_ALPHA_ssub >> art [] \\ + simp [DISJOINT_UNION'] \\ + CONJ_TAC >- rw [Once DISJOINT_SYM] \\ + simp [Abbr ‘t’, Abbr ‘Ns2x’, appstar_APPEND, FV_appstar_MAP_VAR] \\ + reverse CONJ_TAC + >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘set xs2’ >> art [] \\ + simp [Abbr ‘xs2’, LIST_TO_SET_SNOC] \\ + SET_TAC []) \\ + simp [FV_appstar] \\ + CONJ_TAC + >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs2) (set ys2)’ MP_TAC \\ + rw [Abbr ‘xs2’, LIST_TO_SET_SNOC, DISJOINT_ALT]) \\ + simp [MEM_EL] >> rpt STRIP_TAC \\ + Q.PAT_X_ASSUM ‘_ = FV x’ (REWRITE_TAC o wrap) >> POP_ORW \\ + rename1 ‘i < m2’ >> POP_ASSUM MP_TAC \\ + simp [Abbr ‘Ns2''’, EL_MAP] >> DISCH_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET' \\ + Q.EXISTS_TAC ‘FV (EL i Ns2')’ \\ + reverse CONJ_TAC + >- (MP_TAC (Q.SPECL [‘ss’, ‘EL i Ns2'’] FV_ISUB_upperbound) \\ + simp [EL_MAP, Abbr ‘Ns2'’]) \\ + simp [Abbr ‘Ns2'’, EL_listpm, Abbr ‘pm’, REVERSE_ZIP] \\ + (* applying FV_tpm_disjoint *) + ONCE_REWRITE_TAC [DISJOINT_SYM] \\ + MATCH_MP_TAC FV_tpm_disjoint \\ + simp [ALL_DISTINCT_REVERSE] \\ + (* goal: DISJOINT (set ys2) (FV (EL i Ns2)) *) + Know ‘FV N0' SUBSET X UNION RANK r1’ + >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV N'’ >> art [] \\ + qunabbrev_tac ‘N0'’ \\ + MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) >> DISCH_TAC \\ + (* applying FV_subterm_lemma *) + Know ‘FV (EL i Ns2) SUBSET FV N' UNION set vs2’ + >- (MATCH_MP_TAC FV_subterm_lemma \\ + qexistsl_tac [‘X’, ‘r1’, ‘N0'’, ‘n2’, ‘m2’, ‘N1'’] >> simp []) \\ + DISCH_TAC \\ + MATCH_MP_TAC DISJOINT_SUBSET \\ + Q.EXISTS_TAC ‘FV N' UNION set vs2’ >> art [] \\ + REWRITE_TAC [DISJOINT_UNION'] \\ + reverse CONJ_TAC + >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs4’ MP_TAC \\ + ‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) \\ MATCH_MP_TAC DISJOINT_SUBSET \\ Q.EXISTS_TAC ‘X UNION RANK r1’ >> art [] \\ MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set vs3’ \\ + Q.EXISTS_TAC ‘set vs4’ \\ reverse CONJ_TAC - >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + >- (‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ simp [SUBSET_DEF]) \\ simp [DISJOINT_UNION'] \\ - qunabbrev_tac ‘vs3’ \\ + qunabbrev_tac ‘vs4’ \\ MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> Rewr' \\ - ‘FDOM pm1 = set xs1’ by simp [Abbr ‘pm1’, FDOM_fromPairs] \\ - ‘MEM h xs1’ by simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] \\ + ‘vs2 ++ ys2 = vs4’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + ‘FDOM pm2 = set xs2’ by simp [Abbr ‘pm2’, FDOM_fromPairs] \\ + ‘MEM h xs2’ by simp [Abbr ‘xs2’, LIST_TO_SET_SNOC] \\ simp [Abbr ‘t’, ssub_appstar] \\ - Know ‘pm1 ' h = VAR (LAST ys1)’ - >- (‘h = LAST xs1’ by rw [Abbr ‘xs1’, LAST_SNOC] >> POP_ORW \\ - ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ - ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ - simp [Abbr ‘pm1’, LAST_EL] \\ - qabbrev_tac ‘j4 = PRE (LENGTH ys1)’ \\ - ‘0 < LENGTH ys1’ by rw [LENGTH_NON_NIL] \\ - ‘j4 < LENGTH ys1’ by rw [Abbr ‘j4’] \\ - ‘VAR (EL j4 ys1) = EL j4 (MAP VAR ys1)’ by simp [EL_MAP] >> POP_ORW \\ + Know ‘pm2 ' h = VAR (LAST ys2)’ + >- (‘h = LAST xs2’ by rw [Abbr ‘xs2’, LAST_SNOC] >> POP_ORW \\ + ‘xs2 <> []’ by simp [Abbr ‘xs2’] \\ + ‘ys2 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + simp [Abbr ‘pm2’, LAST_EL] \\ + qabbrev_tac ‘j5 = PRE (LENGTH ys2)’ \\ + ‘0 < LENGTH ys2’ by rw [LENGTH_NON_NIL] \\ + ‘j5 < LENGTH ys2’ by rw [Abbr ‘j5’] \\ + ‘VAR (EL j5 ys2) = EL j5 (MAP VAR ys2)’ by simp [EL_MAP] >> POP_ORW \\ MATCH_MP_TAC fromPairs_FAPPLY_EL >> simp []) >> Rewr' \\ - Q.PAT_X_ASSUM ‘_ = W1’ (ONCE_REWRITE_TAC o wrap o SYM) \\ - simp [] \\ - Know ‘LAST ys1 = LAST vs3’ - >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ - ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ + Q.PAT_X_ASSUM ‘_ = W1'’ (ONCE_REWRITE_TAC o wrap o SYM) \\ + Know ‘LAST ys2 = LAST vs4’ + >- (‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ + ‘xs2 <> []’ by simp [Abbr ‘xs2’] \\ + ‘ys2 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ rw [LAST_APPEND_NOT_NIL]) >> Rewr' \\ - STRIP_TAC (* y4 = LAST vs4, etc. *) \\ - (* stage work *) - Know ‘n2 <= n_max’ (* n2 = LAMl_size N0' *) - >- (qunabbrev_tac ‘n_max’ \\ - Q_TAC (TRANS_TAC LESS_EQ_TRANS) ‘subterm_length (M j2) p’ \\ - reverse CONJ_TAC - >- (MATCH_MP_TAC MAX_LIST_PROPERTY >> simp [MEM_MAP] \\ - Q.EXISTS_TAC ‘M j2’ >> art [] \\ - simp [Abbr ‘M’, EL_MEM]) \\ - qunabbrevl_tac [‘n2’, ‘N0'’] \\ - ‘N' = subterm' X (M j2) p r’ by rw [] >> POP_ORW \\ - MATCH_MP_TAC subterm_length_last >> simp []) >> DISCH_TAC \\ - Know ‘n2 < n3’ - >- (Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘n_max’ >> art [] \\ - Q.PAT_X_ASSUM ‘SUC d_max + n1 - m1 = n3’ (REWRITE_TAC o wrap o SYM) \\ - Q_TAC (TRANS_TAC LESS_EQ_LESS_TRANS) ‘d_max + n1 - m1’ \\ - simp [Abbr ‘d_max’]) >> DISCH_TAC \\ - (* applying subterm_headvar_lemma' on W *) - Know ‘hnf_headvar W1' IN FV W' UNION set vs4’ - >- (MATCH_MP_TAC subterm_headvar_lemma' \\ - qexistsl_tac [‘X’, ‘r1’, ‘W0'’, ‘n4’] >> simp []) \\ - Know ‘hnf_headvar W1' = y4’ - >- (Q.PAT_X_ASSUM ‘_ = W1'’ (REWRITE_TAC o wrap o SYM) \\ - simp []) >> Rewr' >> art [] (* y3 -> y1' *) \\ - DISCH_TAC \\ - (* NOTE: FV W' SUBSET X UNION RANK r1 *) - qabbrev_tac ‘X' = X UNION RANK r1’ \\ - Know ‘y2' IN X' UNION set vs4’ - >- (Q.PAT_X_ASSUM ‘y2' IN _’ MP_TAC \\ - Q.PAT_X_ASSUM ‘FV W' SUBSET _’ MP_TAC \\ - SET_TAC []) \\ - Q.PAT_X_ASSUM ‘y2' IN _’ K_TAC >> DISCH_TAC \\ - qunabbrev_tac ‘X'’ \\ - ‘0 < n3 /\ vs3 <> []’ by simp [GSYM LENGTH_NON_NIL] \\ - simp [LAST_EL] \\ - Q.PAT_X_ASSUM ‘y2' IN _’ MP_TAC >> simp [IN_UNION] \\ - STRIP_TAC >| (* 3 subgoals *) - [ (* goal 1 (of 3) *) - SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ - Know ‘MEM y2' vs3’ - >- (POP_ORW >> simp [EL_MEM]) >> DISCH_TAC \\ - Q.PAT_X_ASSUM ‘DISJOINT (set vs3) X’ MP_TAC \\ - simp [DISJOINT_ALT'] >> Q.EXISTS_TAC ‘y2'’ >> art [], - (* goal 2 (of 3) *) - SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ - Know ‘MEM y2' vs3’ - >- (POP_ORW >> simp [EL_MEM]) >> DISCH_TAC \\ - Know ‘DISJOINT (set vs3) (RANK r1)’ - >- simp [Abbr ‘vs3’, DISJOINT_RNEWS_RANK'] \\ - simp [DISJOINT_ALT] \\ - Q.EXISTS_TAC ‘y2'’ >> art [], - (* goal 3 (of 3) *) - Know ‘vs4 <<= vs3’ - >- (qunabbrevl_tac [‘vs3’, ‘vs4’] \\ - MATCH_MP_TAC RNEWS_prefix >> simp []) \\ - simp [IS_PREFIX_APPEND] \\ - DISCH_THEN (Q.X_CHOOSE_THEN ‘ls’ STRIP_ASSUME_TAC) \\ - Know ‘LENGTH ls = n3 - n2’ - >- (POP_ASSUM (MP_TAC o AP_TERM “LENGTH :string list -> num”) \\ - simp []) >> DISCH_TAC \\ - ‘ls <> []’ by simp [GSYM LENGTH_NON_NIL] \\ - Know ‘EL (PRE n3) vs3 = LAST ls’ - >- (Q.PAT_X_ASSUM ‘vs3 = vs4 ++ ls’ (REWRITE_TAC o wrap) \\ - simp [EL_APPEND2, LAST_EL, PRE_SUB]) >> Rewr' \\ - SPOSE_NOT_THEN (ASSUME_TAC o REWRITE_RULE [Once EQ_SYM_EQ]) \\ - Know ‘MEM y2' ls’ - >- (POP_ORW >> simp [MEM_LAST_NOT_NIL]) >> DISCH_TAC \\ - Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ - Q.PAT_X_ASSUM ‘vs3 = vs4 ++ ls’ (REWRITE_TAC o wrap) \\ - simp [ALL_DISTINCT_APPEND] \\ - DISJ2_TAC >> Q.EXISTS_TAC ‘y2'’ >> art [] ]) - (* Case 4 (of 4): hardest *) - >> NTAC 2 (POP_ASSUM (MP_TAC o REWRITE_RULE [])) - >> simp [] - >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j3’ STRIP_ASSUME_TAC) - >> DISCH_THEN (Q.X_CHOOSE_THEN ‘j4’ STRIP_ASSUME_TAC) - >> qabbrev_tac ‘X' = set vs3 UNION set vs4 UNION - FV W1 UNION FV W1' UNION - BIGUNION (IMAGE FV (set Ns1'')) UNION - BIGUNION (IMAGE FV (set Ns2''))’ - >> ‘FINITE X'’ by rw [Abbr ‘X'’] - >> qabbrev_tac ‘d' = MAX (MAX n3 n4) (SUC d_max)’ - >> Q_TAC (NEWS_TAC (“L :string list”, “d' :num”)) ‘X'’ - >> ‘SUC d_max <= LENGTH L /\ n3 <= LENGTH L /\ n4 <= LENGTH L’ - by simp [Abbr ‘d'’, MAX_LE] - >> Know ‘DISJOINT (set L) (set vs1) /\ - DISJOINT (set L) (set vs2) /\ - DISJOINT (set L) (set vs3) /\ - DISJOINT (set L) (set vs4)’ - >- (rw [Abbr ‘L’, Abbr ‘vs1’, Abbr ‘vs2’, Abbr ‘vs3’, Abbr ‘vs4’] \\ - MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) - >> STRIP_TAC - >> Q.PAT_X_ASSUM ‘FINITE X'’ K_TAC - >> Q.PAT_X_ASSUM ‘DISJOINT (set L) X'’ MP_TAC - >> qunabbrev_tac ‘X'’ - >> DISCH_THEN (STRIP_ASSUME_TAC o REWRITE_RULE [DISJOINT_UNION']) - >> STRIP_TAC (* W -h->* ... *) - (* applying hreduce_permutator_shared *) - >> MP_TAC (Q.SPECL [‘Ns1''’, ‘d_max’, ‘L’] hreduce_permutator_shared) - >> simp [] - >> DISCH_THEN (Q.X_CHOOSE_THEN ‘zs1’ (Q.X_CHOOSE_THEN ‘z1’ STRIP_ASSUME_TAC)) - (* applying hreduce_permutator_shared again *) - >> MP_TAC (Q.SPECL [‘Ns2''’, ‘d_max’, ‘L’] hreduce_permutator_shared) - >> simp [] - >> DISCH_THEN (Q.X_CHOOSE_THEN ‘zs2’ (Q.X_CHOOSE_THEN ‘z2’ STRIP_ASSUME_TAC)) - >> Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ MP_TAC - >> Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ MP_TAC - >> qabbrev_tac ‘h = LAST L’ (* The new shared head variable *) - >> qabbrev_tac ‘L' = FRONT L’ - >> ‘L <> []’ by rw [GSYM LENGTH_NON_NIL] - >> NTAC 2 (Q.PAT_X_ASSUM ‘IS_SUFFIX L _’ MP_TAC) - >> ‘L = SNOC h L'’ by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] - >> POP_ORW >> simp [IS_SUFFIX] - >> NTAC 2 STRIP_TAC - >> Q.PAT_X_ASSUM ‘z1 = z2’ (simp o wrap o SYM) - >> Q.PAT_X_ASSUM ‘h = z1’ (simp o wrap o SYM) - >> NTAC 2 DISCH_TAC - >> qabbrev_tac ‘xs1 = SNOC h zs1’ (* suffix of L *) - >> qabbrev_tac ‘xs2 = SNOC h zs2’ (* suffix of L *) - >> Know ‘IS_SUFFIX L xs1 /\ IS_SUFFIX L xs2’ - >- (‘L = SNOC h L'’ - by ASM_SIMP_TAC std_ss [Abbr ‘L'’, Abbr ‘h’, SNOC_LAST_FRONT] \\ - POP_ORW >> simp [IS_SUFFIX, Abbr ‘xs1’, Abbr ‘xs2’]) - >> STRIP_TAC - >> ‘ALL_DISTINCT xs1 /\ ALL_DISTINCT xs2’ by PROVE_TAC [IS_SUFFIX_ALL_DISTINCT] - >> Know ‘LAMl vs1 (P @* Ns1'') -h->* - LAMl vs1 (LAMl zs1 (LAM h (VAR h @* Ns1'' @* MAP VAR zs1))) /\ - LAMl vs2 (P @* Ns2'') -h->* - LAMl vs2 (LAMl zs2 (LAM h (VAR h @* Ns2'' @* MAP VAR zs2)))’ - >- simp [hreduce_LAMl] - >> Q.PAT_X_ASSUM ‘P @* Ns1'' -h->* _’ K_TAC - >> Q.PAT_X_ASSUM ‘P @* Ns2'' -h->* _’ K_TAC - >> REWRITE_TAC [GSYM LAMl_APPEND, GSYM appstar_APPEND] - >> qabbrev_tac ‘Ns1x = Ns1'' ++ MAP VAR zs1’ - >> qabbrev_tac ‘Ns2x = Ns2'' ++ MAP VAR zs2’ - >> REWRITE_TAC [GSYM LAMl_SNOC] - >> qabbrev_tac ‘zs1' = SNOC h (vs1 ++ zs1)’ - >> qabbrev_tac ‘zs2' = SNOC h (vs2 ++ zs2)’ - >> STRIP_TAC - >> Know ‘W -h->* LAMl zs1' (VAR h @* Ns1x) /\ - W' -h->* LAMl zs2' (VAR h @* Ns2x)’ - >- PROVE_TAC [hreduce_TRANS] - >> NTAC 2 (POP_ASSUM K_TAC) >> STRIP_TAC - >> Know ‘LAMl zs1' (VAR h @* Ns1x) = W0 /\ - LAMl zs2' (VAR h @* Ns2x) = W0'’ - >- (‘hnf (LAMl zs1' (VAR h @* Ns1x)) /\ hnf (LAMl zs2' (VAR h @* Ns2x))’ - by rw [hnf_appstar] \\ - METIS_TAC [principle_hnf_thm']) - >> STRIP_TAC - >> Know ‘LENGTH zs1' = n3 /\ LENGTH zs2' = n4’ - >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘_ = W0’ (REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘_ = W0'’ (REWRITE_TAC o wrap o SYM) >> simp []) - >> STRIP_TAC - (* Current situation: - - N -h->* LAMl vs1 (VAR y1 @* Ns1) (= N0) - N' -h->* LAMl vs2 (VAR y2 @* Ns2) (= N0') - -------------------------------------------- - W -h->* LAMl vs3 (VAR y3 @* Ns3) (= W0) ---+ (y3 = LAST vs3) - W -h->* LAMl vs1 (P @* Ns1'') |= - -h->* LAMl zs1' (VAR h @* Ns1x) ---------+ - -------------------------------------------- - W' -h->* LAMl vs4 (VAR y4 @* Ns4) (= W0') --+ (y4 = LAST vs4) - W' -h->* LAMl vs2 (P @* Ns2'') |= - -h->* LAMl zs2' (VAR h @* Ns2x) ---------+ - - Structure of W0: - - LAMl |<---(vs1)--- vs3 -------(ys1)------->| VAR y3 (= LAST vs3) - LAMl |<----------- zs1' ------------------>| VAR h - LAMl |<----vs1----->|<----zs1----------->|h| VAR h - |<------- xs1 ------->| - n3 = n1 + d_max - m1 +1 - (m3 = m1 + d_max - m1 = d_max) - - Structure of W0': - - LAMl |<---(vs2)----- vs4 ----(ys2)--->| VAR y4 (= LAST vs4) - LAMl |<------------- zs2' ----------->| VAR h - LAMl |<----vs2------>|<----zs2----->|h| VAR h - |<---- xs2 ----->| - n4 = n2 + d_max - m2 +1 - (m4 = m2 + d_max - m2 = d_max) - - NOTE: ‘m2 + n1 <> m1 + n2’ indicates that n3 <> n4, thus vs3 <> vs4, - but they are both the prefix of a longer list. Thus y3 and y4 cannot be - the same, because they are LAST vs3 and LAST vs4, different positions - of this long list, which is ALL_DISTINCT. - *) - >> DISCH_TAC >> DISJ1_TAC (* the second part is actually equivalent *) - >> Know ‘n3 <> n4’ - >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ - simp [Abbr ‘zs1'’, Abbr ‘zs2'’]) - >> DISCH_TAC - (* properties of n3 and n4 *) - >> Know ‘SUC d_max + n1 - m1 = n3’ - >- (Q.PAT_X_ASSUM ‘_ = n3’ (REWRITE_TAC o wrap o SYM) \\ - simp [Abbr ‘zs1'’]) - >> DISCH_TAC - >> Know ‘SUC d_max + n2 - m2 = n4’ - >- (Q.PAT_X_ASSUM ‘_ = n4’ (REWRITE_TAC o wrap o SYM) \\ - simp [Abbr ‘zs2'’]) + simp []) >> DISCH_TAC - (* stage work *) - >> Suff ‘y3 = LAST vs3 /\ y4 = LAST vs4’ - >- (qabbrev_tac ‘n5 = MAX n3 n4’ \\ + >> PRINT_TAC "stage work on subtree_equiv_lemma: L9905" + >> Know ‘y3 = y4 <=> n3 = n4’ + >- (Q.PAT_X_ASSUM ‘LAST vs3 = y3’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘LAST vs4 = y4’ (REWRITE_TAC o wrap o SYM) \\ + qabbrev_tac ‘n5 = MAX n3 n4’ \\ Q_TAC (RNEWS_TAC (“vs5 :string list”, “r1 :num”, “n5 :num”)) ‘X’ \\ ‘n3 <= n5 /\ n4 <= n5’ by simp [Abbr ‘n5’, MAX_LE] \\ Know ‘vs3 <<= vs5’ @@ -9233,251 +9919,56 @@ Proof Q.PAT_X_ASSUM ‘vs4 = TAKE n4' vs5’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ POP_ORW >> NTAC 2 DISCH_TAC \\ ‘vs3 <> [] /\ vs4 <> []’ by simp [GSYM LENGTH_NON_NIL] \\ - simp [LAST_EL] >> DISCH_THEN K_TAC \\ - Q.PAT_X_ASSUM ‘_ = vs3’ (REWRITE_TAC o wrap o SYM) \\ - Q.PAT_X_ASSUM ‘_ = vs4’ (REWRITE_TAC o wrap o SYM) \\ + simp [LAST_EL] \\ + Q.PAT_X_ASSUM ‘TAKE n3 vs5 = vs3’ (REWRITE_TAC o wrap o SYM) \\ + Q.PAT_X_ASSUM ‘TAKE n4 vs5 = vs4’ (REWRITE_TAC o wrap o SYM) \\ simp [EL_TAKE] \\ (* applying ALL_DISTINCT_EL_IMP *) Know ‘EL (PRE n3) vs5 = EL (PRE n4) vs5 <=> PRE n3 = PRE n4’ >- (MATCH_MP_TAC ALL_DISTINCT_EL_IMP >> simp []) >> Rewr' \\ simp [INV_PRE_EQ]) - (* now proving ‘y3 = LAST vs3 /\ y4 = LAST vs4’ *) - >> CONJ_TAC - >| [ (* goal 1 (of 2) *) - Know ‘vs1 <<= vs3’ - >- (qunabbrevl_tac [‘vs1’, ‘vs3’] \\ - MATCH_MP_TAC RNEWS_prefix >> simp []) \\ - simp [IS_PREFIX_EQ_TAKE] \\ - DISCH_THEN (Q.X_CHOOSE_THEN ‘n1'’ STRIP_ASSUME_TAC) \\ - ‘n1' = n1’ by rw [] \\ - Q.PAT_X_ASSUM ‘n1' <= n3’ MP_TAC \\ - Q.PAT_X_ASSUM ‘vs1 = TAKE n1' vs3’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ - POP_ORW >> NTAC 2 DISCH_TAC \\ - ‘zs1' = vs1 ++ xs1’ by simp [Abbr ‘zs1'’, Abbr ‘xs1’] \\ - qabbrev_tac ‘ys1 = DROP n1 vs3’ \\ - ‘ALL_DISTINCT ys1’ by simp [Abbr ‘ys1’, ALL_DISTINCT_DROP] \\ - Know ‘DISJOINT (set xs1) (set ys1)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs3’ \\ - reverse CONJ_TAC >- simp [LIST_TO_SET_DROP, Abbr ‘ys1’] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set L’ >> art [] \\ - MATCH_MP_TAC LIST_TO_SET_SUFFIX >> art []) >> DISCH_TAC \\ - ‘LENGTH xs1 = LENGTH ys1’ by simp [Abbr ‘xs1’, Abbr ‘ys1’] \\ - Know ‘LAMl vs3 (VAR y3 @* Ns3) = LAMl zs1' (VAR h @* Ns1x)’ >- rw [] \\ - ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - Q.PAT_X_ASSUM ‘zs1' = _’ (ONCE_REWRITE_TAC o wrap) \\ - simp [LAMl_APPEND] \\ - qabbrev_tac ‘t = VAR h @* Ns1x’ \\ - (* applying LAMl_ALPHA_ssub *) - qabbrev_tac ‘pm1 = fromPairs xs1 (MAP VAR ys1)’ \\ - (* NOTE: The following disjointness hold for names from different rows *) - Know ‘DISJOINT (set vs) (set ys1)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘vs’, ‘vs3’] \\ - MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ - Know ‘DISJOINT (set ys) (set ys1)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs3’ >> simp [Abbr ‘ys1’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘ys’, ‘vs3’] \\ - MATCH_MP_TAC DISJOINT_RNEWS \\ - ‘0 < LENGTH p’ by rw [LENGTH_NON_NIL] \\ - simp [Abbr ‘r1’]) >> DISCH_TAC \\ - Know ‘LAMl xs1 t = LAMl ys1 (pm1 ' t)’ - >- (simp [Abbr ‘pm1’, fromPairs_def] \\ - MATCH_MP_TAC LAMl_ALPHA_ssub >> art [] \\ - simp [DISJOINT_UNION'] \\ - CONJ_TAC >- rw [Once DISJOINT_SYM] \\ - simp [Abbr ‘t’, Abbr ‘Ns1x’, appstar_APPEND, FV_appstar_MAP_VAR] \\ - reverse CONJ_TAC - >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set xs1’ >> art [] \\ - simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] >> SET_TAC []) \\ - simp [FV_appstar] \\ - CONJ_TAC - >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs1) (set ys1)’ MP_TAC \\ - rw [Abbr ‘xs1’, LIST_TO_SET_SNOC, DISJOINT_ALT]) \\ - simp [MEM_EL] >> rpt STRIP_TAC \\ - Q.PAT_X_ASSUM ‘_ = FV x’ (REWRITE_TAC o wrap) >> POP_ORW \\ - rename1 ‘i < m1’ >> POP_ASSUM MP_TAC \\ - simp [Abbr ‘Ns1''’, EL_MAP] >> DISCH_TAC \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘FV (EL i Ns1')’ \\ - reverse CONJ_TAC - >- (MP_TAC (Q.SPECL [‘ss’, ‘EL i Ns1'’] FV_ISUB_upperbound) \\ - simp [EL_MAP, Abbr ‘Ns1'’]) \\ - simp [Abbr ‘Ns1'’, EL_listpm, Abbr ‘pm’, REVERSE_ZIP] \\ - (* applying FV_tpm_disjoint *) - ONCE_REWRITE_TAC [DISJOINT_SYM] \\ - MATCH_MP_TAC FV_tpm_disjoint \\ - simp [ALL_DISTINCT_REVERSE] \\ - (* goal: DISJOINT (set ys1) (FV (EL i Ns1)) *) - Know ‘FV N0 SUBSET X UNION RANK r1’ - >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV N’ >> art [] \\ - qunabbrev_tac ‘N0’ \\ - MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) >> DISCH_TAC \\ - (* applying FV_subterm_lemma *) - Know ‘FV (EL i Ns1) SUBSET FV N UNION set vs1’ - >- (MATCH_MP_TAC FV_subterm_lemma \\ - qexistsl_tac [‘X’, ‘r1’, ‘N0’, ‘n1’, ‘m1’, ‘N1’] >> simp []) \\ - DISCH_TAC \\ - MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘FV N UNION set vs1’ >> art [] \\ - REWRITE_TAC [DISJOINT_UNION'] \\ - reverse CONJ_TAC - >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs3’ MP_TAC \\ - ‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) \\ - MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘X UNION RANK r1’ >> art [] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set vs3’ \\ - reverse CONJ_TAC - >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - simp [SUBSET_DEF]) \\ - simp [DISJOINT_UNION'] \\ - qunabbrev_tac ‘vs3’ \\ - MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> Rewr' \\ - ‘vs1 ++ ys1 = vs3’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - ‘FDOM pm1 = set xs1’ by simp [Abbr ‘pm1’, FDOM_fromPairs] \\ - ‘MEM h xs1’ by simp [Abbr ‘xs1’, LIST_TO_SET_SNOC] \\ - simp [Abbr ‘t’, ssub_appstar] \\ - Know ‘pm1 ' h = VAR (LAST ys1)’ - >- (‘h = LAST xs1’ by rw [Abbr ‘xs1’, LAST_SNOC] >> POP_ORW \\ - ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ - ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ - simp [Abbr ‘pm1’, LAST_EL] \\ - qabbrev_tac ‘j5 = PRE (LENGTH ys1)’ \\ - ‘0 < LENGTH ys1’ by rw [LENGTH_NON_NIL] \\ - ‘j5 < LENGTH ys1’ by rw [Abbr ‘j5’] \\ - ‘VAR (EL j5 ys1) = EL j5 (MAP VAR ys1)’ by simp [EL_MAP] >> POP_ORW \\ - MATCH_MP_TAC fromPairs_FAPPLY_EL >> simp []) >> Rewr' \\ - Q.PAT_X_ASSUM ‘_ = W1’ (ONCE_REWRITE_TAC o wrap o SYM) \\ - Know ‘LAST ys1 = LAST vs3’ - >- (‘vs3 = vs1 ++ ys1’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - ‘xs1 <> []’ by simp [Abbr ‘xs1’] \\ - ‘ys1 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ - rw [LAST_APPEND_NOT_NIL]) >> Rewr' \\ - simp [], - (* goal 2 (of 2) *) - Know ‘vs2 <<= vs4’ - >- (qunabbrevl_tac [‘vs2’, ‘vs4’] \\ - MATCH_MP_TAC RNEWS_prefix >> simp []) \\ - simp [IS_PREFIX_EQ_TAKE] \\ - DISCH_THEN (Q.X_CHOOSE_THEN ‘n2'’ STRIP_ASSUME_TAC) \\ - ‘n2' = n2’ by rw [] \\ - Q.PAT_X_ASSUM ‘n2' <= n4’ MP_TAC \\ - Q.PAT_X_ASSUM ‘vs2 = TAKE n2' vs4’ (MP_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ]) \\ - POP_ORW >> NTAC 2 DISCH_TAC \\ - ‘zs2' = vs2 ++ xs2’ by simp [Abbr ‘zs2'’, Abbr ‘xs2’] \\ - qabbrev_tac ‘ys2 = DROP n2 vs4’ \\ - ‘ALL_DISTINCT ys2’ by simp [Abbr ‘ys2’, ALL_DISTINCT_DROP] \\ - Know ‘DISJOINT (set xs2) (set ys2)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs4’ \\ - reverse CONJ_TAC >- simp [LIST_TO_SET_DROP, Abbr ‘ys2’] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set L’ >> art [] \\ - MATCH_MP_TAC LIST_TO_SET_SUFFIX >> art []) >> DISCH_TAC \\ - ‘LENGTH xs2 = LENGTH ys2’ by simp [Abbr ‘xs2’, Abbr ‘ys2’] \\ - Know ‘LAMl vs4 (VAR y4 @* Ns4) = LAMl zs2' (VAR h @* Ns2x)’ >- rw [] \\ - ‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - Q.PAT_X_ASSUM ‘zs2' = _’ (ONCE_REWRITE_TAC o wrap) \\ - simp [LAMl_APPEND] \\ - qabbrev_tac ‘t = VAR h @* Ns2x’ \\ - (* applying LAMl_ALPHA_ssub *) - qabbrev_tac ‘pm2 = fromPairs xs2 (MAP VAR ys2)’ \\ - Know ‘DISJOINT (set vs) (set ys2)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘vs’, ‘vs4’] \\ - MATCH_MP_TAC DISJOINT_RNEWS >> simp [Abbr ‘r1’]) >> DISCH_TAC \\ - Know ‘DISJOINT (set ys) (set ys2)’ - >- (MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘set vs4’ >> simp [Abbr ‘ys2’, LIST_TO_SET_DROP] \\ - qunabbrevl_tac [‘ys’, ‘vs4’] \\ - MATCH_MP_TAC DISJOINT_RNEWS \\ - ‘0 < LENGTH p’ by rw [LENGTH_NON_NIL] \\ - simp [Abbr ‘r1’]) >> DISCH_TAC \\ - Know ‘LAMl xs2 t = LAMl ys2 (pm2 ' t)’ - >- (simp [Abbr ‘pm2’, fromPairs_def] \\ - MATCH_MP_TAC LAMl_ALPHA_ssub >> art [] \\ - (* goal: DISJOINT (set ys2) (set xs2 UNION FV t) *) - simp [DISJOINT_UNION'] \\ - CONJ_TAC >- rw [Once DISJOINT_SYM] \\ - simp [Abbr ‘t’, Abbr ‘Ns2x’, appstar_APPEND, FV_appstar_MAP_VAR] \\ - reverse CONJ_TAC - >- (MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set xs2’ >> art [] \\ - simp [Abbr ‘xs2’, LIST_TO_SET_SNOC] \\ - SET_TAC []) \\ - simp [FV_appstar] \\ - CONJ_TAC - >- (Q.PAT_X_ASSUM ‘DISJOINT (set xs2) (set ys2)’ MP_TAC \\ - rw [Abbr ‘xs2’, LIST_TO_SET_SNOC, DISJOINT_ALT]) \\ - simp [MEM_EL] >> rpt STRIP_TAC \\ - Q.PAT_X_ASSUM ‘_ = FV x’ (REWRITE_TAC o wrap) >> POP_ORW \\ - rename1 ‘i < m2’ >> POP_ASSUM MP_TAC \\ - simp [Abbr ‘Ns2''’, EL_MAP] >> DISCH_TAC \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘FV (EL i Ns2')’ \\ - reverse CONJ_TAC - >- (MP_TAC (Q.SPECL [‘ss’, ‘EL i Ns2'’] FV_ISUB_upperbound) \\ - simp [EL_MAP, Abbr ‘Ns2'’]) \\ - simp [Abbr ‘Ns2'’, EL_listpm, Abbr ‘pm’, REVERSE_ZIP] \\ - (* applying FV_tpm_disjoint *) - ONCE_REWRITE_TAC [DISJOINT_SYM] \\ - MATCH_MP_TAC FV_tpm_disjoint \\ - simp [ALL_DISTINCT_REVERSE] \\ - (* goal: DISJOINT (set ys2) (FV (EL i Ns2)) *) - Know ‘FV N0' SUBSET X UNION RANK r1’ - >- (Q_TAC (TRANS_TAC SUBSET_TRANS) ‘FV N'’ >> art [] \\ - qunabbrev_tac ‘N0'’ \\ - MATCH_MP_TAC principle_hnf_FV_SUBSET' >> art []) >> DISCH_TAC \\ - (* applying FV_subterm_lemma *) - Know ‘FV (EL i Ns2) SUBSET FV N' UNION set vs2’ - >- (MATCH_MP_TAC FV_subterm_lemma \\ - qexistsl_tac [‘X’, ‘r1’, ‘N0'’, ‘n2’, ‘m2’, ‘N1'’] >> simp []) \\ - DISCH_TAC \\ - MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘FV N' UNION set vs2’ >> art [] \\ - REWRITE_TAC [DISJOINT_UNION'] \\ - reverse CONJ_TAC - >- (Q.PAT_X_ASSUM ‘ALL_DISTINCT vs4’ MP_TAC \\ - ‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - simp [ALL_DISTINCT_APPEND', Once DISJOINT_SYM]) \\ - MATCH_MP_TAC DISJOINT_SUBSET \\ - Q.EXISTS_TAC ‘X UNION RANK r1’ >> art [] \\ - MATCH_MP_TAC DISJOINT_SUBSET' \\ - Q.EXISTS_TAC ‘set vs4’ \\ - reverse CONJ_TAC - >- (‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - simp [SUBSET_DEF]) \\ - simp [DISJOINT_UNION'] \\ - qunabbrev_tac ‘vs4’ \\ - MATCH_MP_TAC DISJOINT_RNEWS_RANK' >> art []) >> Rewr' \\ - ‘vs2 ++ ys2 = vs4’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - ‘FDOM pm2 = set xs2’ by simp [Abbr ‘pm2’, FDOM_fromPairs] \\ - ‘MEM h xs2’ by simp [Abbr ‘xs2’, LIST_TO_SET_SNOC] \\ - simp [Abbr ‘t’, ssub_appstar] \\ - Know ‘pm2 ' h = VAR (LAST ys2)’ - >- (‘h = LAST xs2’ by rw [Abbr ‘xs2’, LAST_SNOC] >> POP_ORW \\ - ‘xs2 <> []’ by simp [Abbr ‘xs2’] \\ - ‘ys2 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ - simp [Abbr ‘pm2’, LAST_EL] \\ - qabbrev_tac ‘j5 = PRE (LENGTH ys2)’ \\ - ‘0 < LENGTH ys2’ by rw [LENGTH_NON_NIL] \\ - ‘j5 < LENGTH ys2’ by rw [Abbr ‘j5’] \\ - ‘VAR (EL j5 ys2) = EL j5 (MAP VAR ys2)’ by simp [EL_MAP] >> POP_ORW \\ - MATCH_MP_TAC fromPairs_FAPPLY_EL >> simp []) >> Rewr' \\ - Q.PAT_X_ASSUM ‘_ = W1'’ (ONCE_REWRITE_TAC o wrap o SYM) \\ - Know ‘LAST ys2 = LAST vs4’ - >- (‘vs4 = vs2 ++ ys2’ by METIS_TAC [TAKE_DROP] >> POP_ORW \\ - ‘xs2 <> []’ by simp [Abbr ‘xs2’] \\ - ‘ys2 <> []’ by METIS_TAC [LENGTH_NON_NIL] \\ - rw [LAST_APPEND_NOT_NIL]) >> Rewr' \\ - simp [] ] + >> Rewr' + (* Current situation: + + N -h->* LAMl vs1 (VAR y1 @* Ns1) (= N0) + N' -h->* LAMl vs2 (VAR y2 @* Ns2) (= N0') + -------------------------------------------- + W -h->* LAMl vs3 (VAR y3 @* Ns3) (= W0) ---+ (y3 = LAST vs3) + W -h->* LAMl vs1 (P3 @* Ns1'') |= + -h->* LAMl zs1' (VAR h @* Ns1x) ---------+ + -------------------------------------------- + W' -h->* LAMl vs4 (VAR y4 @* Ns4) (= W0') --+ (y4 = LAST vs4) + W' -h->* LAMl vs2 (P4 @* Ns2'') |= + -h->* LAMl zs2' (VAR h @* Ns2x) ---------+ + + Structure of W0: + + LAMl |<---(vs1)--- vs3 -------(ys1)------->| VAR y3 (= LAST vs3) + LAMl |<----------- zs1' ------------------>| VAR h + LAMl |<--- vs1 ---->|<------- zs1 ------>|h| VAR h + |<------- xs1 -------->| + n3 = n1 + d_max' j3 - m1 + 1 + (m3 = m1 + d_max' j3 - m1 = d_max' j3) + + Structure of W0': + + LAMl |<---(vs2)----- vs4 ----(ys2)--->| VAR y4 (= LAST vs4) + LAMl |<------------- zs2' ----------->| VAR h + LAMl |<--- vs2 ----->|<---- zs2 --->|h| VAR h + |<---- xs2 ----->| + n4 = n2 + d_max' j4 - m2 + 1 + (m4 = m2 + d_max' j4 - m2 = d_max' j4) + *) + >> PRINT_TAC "stage work on subtree_equiv_lemma: L9975" + >> Cases_on ‘y1 = y2’ >> simp [] + (* now: y1 <> y2 *) + >> ‘y1' <> y2'’ by rw [Abbr ‘y1'’, Abbr ‘y2'’] + >> ‘y j3 <> y j4’ by rw [] + >> Suff ‘m3 <> m4’ >- simp [] + (* final goal (uniqueness of f) *) + >> Q.PAT_X_ASSUM ‘_ = m3’ (REWRITE_TAC o wrap o SYM) + >> Q.PAT_X_ASSUM ‘_ = m4’ (REWRITE_TAC o wrap o SYM) + >> simp [Abbr ‘d_max'’] QED val _ = export_theory (); diff --git a/examples/lambda/barendregt/head_reductionScript.sml b/examples/lambda/barendregt/head_reductionScript.sml index 6a3a3d3b41..03a1473906 100644 --- a/examples/lambda/barendregt/head_reductionScript.sml +++ b/examples/lambda/barendregt/head_reductionScript.sml @@ -2531,6 +2531,21 @@ Proof >> PROVE_TAC [hreduce_abs] (* is_abs N *) QED +(* NOTE: This theorem depends on LAMl_size and cannot move to chap2Theory *) +Theorem permutator_11[simp] : + permutator m = permutator n <=> m = n +Proof + reverse EQ_TAC >- rw [] + >> rw [permutator_def, GENLIST] + >> qabbrev_tac ‘vs1 = GENLIST n2s m’ + >> qabbrev_tac ‘vs2 = GENLIST n2s n’ + >> ‘LENGTH vs1 = m /\ LENGTH vs2 = n’ by rw [Abbr ‘vs1’, Abbr ‘vs2’] + >> Q.PAT_X_ASSUM ‘LAMl vs1 _ = LAMl vs2 _’ + (MP_TAC o AP_TERM “LAMl_size :term -> num”) + >> REWRITE_TAC [GSYM LAMl_SNOC, LAMl_size_hnf] + >> simp [] +QED + val _ = export_theory() val _ = html_theory "head_reduction"; diff --git a/examples/lambda/barendregt/lameta_completeScript.sml b/examples/lambda/barendregt/lameta_completeScript.sml index c595e3c3e3..c5ff4cc31f 100644 --- a/examples/lambda/barendregt/lameta_completeScript.sml +++ b/examples/lambda/barendregt/lameta_completeScript.sml @@ -7,13 +7,12 @@ open HolKernel Parse boolLib bossLib; -open hurdUtils arithmeticTheory pred_setTheory listTheory rich_listTheory; +open hurdUtils arithmeticTheory pred_setTheory listTheory rich_listTheory + ltreeTheory llistTheory relationTheory; open termTheory basic_swapTheory appFOLDLTheory chap2Theory chap3Theory horeductionTheory solvableTheory head_reductionTheory head_reductionLib - boehmTheory; - -val _ = new_theory "lameta_complete"; + standardisationTheory boehmTheory; val _ = temp_delsimps [ "lift_disj_eq", "lift_imp_disj", @@ -30,24 +29,286 @@ val _ = hide "C"; val _ = hide "W"; val _ = hide "Y"; +val _ = new_theory "lameta_complete"; + +(*---------------------------------------------------------------------------* + * ltreeTheory extras + *---------------------------------------------------------------------------*) + +(* ltree_subset A B <=> A results from B by "cutting off" some subtrees. Thus, + + 1) The paths of A is a subset of paths of B + 2) The node contents for all paths of A is identical to those of B, but the number + of successors at those nodes of B may be different (B may have more successors) + + NOTE: Simply defining ‘ltree_subset A B <=> subtrees A SUBSET subtrees B’ is wrong, + as A may appear as a non-root subtree of B, i.e. ‘A IN subtrees B’ but there's + no way to "cut off" B to get A, the resulting subtree in B always have some + more path prefixes. + *) +Definition ltree_subset_def : + ltree_subset A B <=> + (ltree_paths A) SUBSET (ltree_paths B) /\ + !p. p IN ltree_paths A ==> + ltree_node (THE (ltree_lookup A p)) = + ltree_node (THE (ltree_lookup B p)) +End + +(*---------------------------------------------------------------------------* + * Equivalent terms + *---------------------------------------------------------------------------*) + +(* Definition 10.2.19 [1, p. 238] + + M = LAMl v1 (y @* Ms) @@ (MAP VAR v1) == y @* Ms' + N = LAMl v2 (y' @* Ns) @@ (MAP VAR v2) == y' @* Ns' + + LENGTH Ms = n /\ LENGTH Ns = n' + LENGTH Ms' = m /\ LENGTH Ns' = m' + + y = y' + n - m = n' - m' (possibly negative) <=> n + m' = n' + m (non-negative) + *) +Definition equivalent_def : + equivalent (M :term) (N :term) = + if solvable M /\ solvable N then + let M0 = principle_hnf M; + N0 = principle_hnf N; + n = LAMl_size M0; + n' = LAMl_size N0; + vs = NEWS (MAX n n') (FV M UNION FV N); + vsM = TAKE n vs; + vsN = TAKE n' vs; + M1 = principle_hnf (M0 @* MAP VAR vsM); + N1 = principle_hnf (N0 @* MAP VAR vsN); + y = hnf_head M1; + y' = hnf_head N1; + m = LENGTH (hnf_children M1); + m' = LENGTH (hnf_children N1); + in + y = y' /\ n + m' = n' + m + else + ~solvable M /\ ~solvable N +End + +Theorem equivalent_reflexive : + reflexive equivalent +Proof + rw [reflexive_def, equivalent_def] +QED + +(* |- equivalent x x *) +Theorem equivalent_refl[simp] = + SPEC_ALL (REWRITE_RULE [reflexive_def] equivalent_reflexive) + +Theorem equivalent_symmetric : + symmetric equivalent +Proof + RW_TAC std_ss [symmetric_def, equivalent_def, Once MAX_COMM, Once UNION_COMM] + >> reverse (Cases_on ‘solvable x /\ solvable y’) >- fs [] + >> simp [] + >> rename1 ‘y1 = y2 /\ m1 + n = m + n1 <=> y3 = y4 /\ m3 + n1 = m2 + n3’ + >> ‘n3 = n’ by rw [Abbr ‘n3’, Abbr ‘n’] >> gs [] + >> EQ_TAC >> rw [] +QED + +(* |- !x y. equivalent x y <=> equivalent y x *) +Theorem equivalent_comm = REWRITE_RULE [symmetric_def] equivalent_symmetric + +Theorem equivalent_of_solvables : + !M N. solvable M /\ solvable N ==> + (equivalent M N <=> + let M0 = principle_hnf M; + N0 = principle_hnf N; + n = LAMl_size M0; + n' = LAMl_size N0; + vs = NEWS (MAX n n') (FV M UNION FV N); + vsM = TAKE n vs; + vsN = TAKE n' vs; + M1 = principle_hnf (M0 @* MAP VAR vsM); + N1 = principle_hnf (N0 @* MAP VAR vsN); + y = hnf_head M1; + y' = hnf_head N1; + m = LENGTH (hnf_children M1); + m' = LENGTH (hnf_children N1); + in + y = y' /\ n + m' = n' + m) +Proof + RW_TAC std_ss [equivalent_def] +QED + +(* beta-equivalent terms are also equivalent here *) +Theorem lameq_imp_equivalent : + !M N. M == N ==> equivalent M N +Proof + rpt STRIP_TAC + >> reverse (Cases_on ‘solvable M’) + >- (‘unsolvable N’ by METIS_TAC [lameq_solvable_cong] \\ + rw [equivalent_def]) + >> ‘solvable N’ by METIS_TAC [lameq_solvable_cong] + >> qabbrev_tac ‘X = FV M UNION FV N’ + >> ‘FINITE X’ by rw [Abbr ‘X’] + >> ‘LAMl_size (principle_hnf M) = LAMl_size (principle_hnf N)’ + by METIS_TAC [lameq_principle_hnf_size_eq'] + (* stage work *) + >> RW_TAC std_ss [equivalent_of_solvables] (* 2 subgoals, same tactics *) + >> ‘ALL_DISTINCT vs /\ DISJOINT (set vs) X /\ LENGTH vs = n’ + by (rw [Abbr ‘vs’, NEWS_def]) + >> ‘vsM = vs’ by rw [Abbr ‘vsM’, TAKE_LENGTH_ID_rwt] + >> POP_ASSUM (fs o wrap) + >> Q.PAT_X_ASSUM ‘vs = vsN’ (fs o wrap o SYM) + >> MP_TAC (Q.SPECL [‘X’, ‘M’, ‘N’, ‘M0’, ‘N0’, ‘n’, ‘vs’, ‘M1’, ‘N1’] + lameq_principle_hnf_thm_simple) + >> simp [Abbr ‘X’, GSYM solvable_iff_has_hnf] +QED + +(* NOTE: the initial calls of ‘principle_hnf’ get eliminated if the involved + terms are already in head normal forms. + *) +Theorem equivalent_of_hnf : + !M N. hnf M /\ hnf N ==> + (equivalent M N <=> + let n = LAMl_size M; + n' = LAMl_size N; + vs = NEWS (MAX n n') (FV M UNION FV N); + vsM = TAKE n vs; + vsN = TAKE n' vs; + M1 = principle_hnf (M @* MAP VAR vsM); + N1 = principle_hnf (N @* MAP VAR vsN); + y = hnf_head M1; + y' = hnf_head N1; + m = LENGTH (hnf_children M1); + m' = LENGTH (hnf_children N1) + in + y = y' /\ n + m' = n' + m) +Proof + rpt STRIP_TAC + >> ‘solvable M /\ solvable N’ by PROVE_TAC [hnf_has_hnf, solvable_iff_has_hnf] + >> RW_TAC std_ss [equivalent_def, principle_hnf_reduce] + >> METIS_TAC [] +QED + +(* From [1, p.238]. This concerte example shows that dB encoding is not easy in + defining this "concept": the literal encoding of inner head variables are not + the same for equivalent terms. + *) +Theorem not_equivalent_example : + !x y. x <> y ==> ~equivalent (LAM x (VAR y @@ t)) (LAM y (VAR y @@ t)) +Proof + qx_genl_tac [‘x’, ‘v’] >> DISCH_TAC + >> ‘hnf (LAM x (VAR v @@ t)) /\ hnf (LAM v (VAR v @@ t))’ by rw [] + >> ‘solvable (LAM x (VAR v @@ t)) /\ solvable (LAM v (VAR v @@ t))’ + by rw [solvable_iff_has_hnf, hnf_has_hnf] + >> RW_TAC std_ss [equivalent_of_solvables, principle_hnf_reduce] + (* fix M0 *) + >> qunabbrev_tac ‘M0’ >> qabbrev_tac ‘M0 = LAM x (VAR v @@ t)’ + >> ‘ALL_DISTINCT vs /\ DISJOINT (set vs) (FV M0 UNION FV N0) /\ + LENGTH vs = MAX n n'’ by rw [Abbr ‘vs’, NEWS_def] + >> ‘DISJOINT (set vs) (FV M0) /\ DISJOINT (set vs) (FV N0)’ + by METIS_TAC [DISJOINT_SYM, DISJOINT_UNION] + >> Q_TAC (HNF_TAC (“M0 :term”, “vs :string list”, + “y1 :string”, “args1 :term list”)) ‘M1’ + >> Q_TAC (HNF_TAC (“N0 :term”, “vs :string list”, + “y2 :string”, “args2 :term list”)) ‘N1’ + >> ‘TAKE (LAMl_size M0) vs = vsM’ by rw [Abbr ‘vsM’, Abbr ‘n’] + >> ‘TAKE (LAMl_size N0) vs = vsN’ by rw [Abbr ‘vsN’, Abbr ‘n'’] + >> NTAC 2 (POP_ASSUM (rfs o wrap)) + (* reshaping and reordering assumptions *) + >> qunabbrev_tac ‘M1’ + >> qabbrev_tac ‘M1 = principle_hnf (M0 @* MAP VAR vsM)’ + >> qunabbrev_tac ‘N1’ + >> qabbrev_tac ‘N1 = principle_hnf (N0 @* MAP VAR vsN)’ + >> Q.PAT_X_ASSUM ‘M0 = _’ ASSUME_TAC + >> Q.PAT_X_ASSUM ‘N0 = _’ ASSUME_TAC + >> Q.PAT_X_ASSUM ‘M1 = _’ ASSUME_TAC + >> Q.PAT_X_ASSUM ‘N1 = _’ ASSUME_TAC + >> ‘VAR y1 = y’ by rw [Abbr ‘y’ , absfree_hnf_head] + >> ‘VAR y2 = y'’ by rw [Abbr ‘y'’, absfree_hnf_head] + >> qunabbrevl_tac [‘n’, ‘n'’] + >> Know ‘LAMl_size M0 = 1 /\ LAMl_size N0 = 1’ + >- (rw [Abbr ‘M0’, Abbr ‘N0’, LAMl_size_def]) + >> DISCH_THEN (rfs o wrap) + >> ‘vsN = vs’ by rw [Abbr ‘vsN’, TAKE_LENGTH_ID_rwt] + >> POP_ASSUM (rfs o wrap) + >> Q.PAT_X_ASSUM ‘vs = vsM’ (rfs o wrap o SYM) + >> qunabbrev_tac ‘vsN’ + (* stage work *) + >> qabbrev_tac ‘z = HD vs’ + >> ‘vs = [z]’ by METIS_TAC [SING_HD] + >> POP_ASSUM (rfs o wrap) + >> qunabbrevl_tac [‘M0’, ‘N0’] + >> DISJ1_TAC + >> qunabbrevl_tac [‘y’, ‘y'’] + >> Q.PAT_X_ASSUM ‘VAR y1 = hnf_head M1’ (rfs o wrap o SYM) + >> Q.PAT_X_ASSUM ‘_ = LAM z (VAR y1 @* args1)’ (rfs o wrap o SYM) + >> Q.PAT_X_ASSUM ‘_ = LAM z (VAR y2 @* args2)’ (rfs o wrap o SYM) + (* now the goal is ‘y1 <> y2’ *) + >> qabbrev_tac ‘u = VAR v @@ t’ + >> ‘hnf u’ by rw [Abbr ‘u’] + >> Know ‘M1 = [VAR z/x] u’ + >- (qunabbrev_tac ‘M1’ \\ + Cases_on ‘z = x’ >- (POP_ASSUM (gs o wrap) \\ + fs [principle_hnf_beta_reduce1]) \\ + MATCH_MP_TAC principle_hnf_beta >> simp [Abbr ‘u’] \\ + rfs [FV_thm]) + >> DISCH_THEN (rfs o wrap) + >> Know ‘N1 = [VAR z/v] u’ + >- (qunabbrev_tac ‘N1’ \\ + Cases_on ‘z = v’ >- (POP_ASSUM (rfs o wrap)) \\ + MATCH_MP_TAC principle_hnf_beta >> simp [Abbr ‘u’] \\ + rfs [FV_thm]) + >> DISCH_THEN (rfs o wrap) + >> qunabbrevl_tac [‘M1’, ‘N1’] + >> rfs [Abbr ‘u’, app_eq_appstar] + >> METIS_TAC [] +QED + +Theorem equivalent_of_unsolvables : + !M N. unsolvable M /\ unsolvable N ==> equivalent M N +Proof + rw [equivalent_def] +QED + +(*---------------------------------------------------------------------------* + * subtree_equiv_lemma + *---------------------------------------------------------------------------*) + +Theorem subtree_equiv_lemma_explicit'[local] = + subtree_equiv_lemma_explicit |> SIMP_RULE std_ss [LET_DEF] + +Theorem subtree_equiv_lemma : + !X Ms p r. + FINITE X /\ p <> [] /\ 0 < r /\ Ms <> [] /\ + BIGUNION (IMAGE FV (set Ms)) SUBSET X UNION RANK r /\ + EVERY (\M. subterm X M p r <> NONE) Ms ==> + ?pi. Boehm_transform pi /\ EVERY is_ready (MAP (apply pi) Ms) /\ + EVERY (\M. p IN ltree_paths (BT' X M r)) (MAP (apply pi) Ms) /\ + !M N q. MEM M Ms /\ MEM N Ms /\ q <<= p ==> + (subtree_equiv X M N q r <=> + subtree_equiv X (apply pi M) (apply pi N) q r) +Proof + rpt STRIP_TAC + >> Q.EXISTS_TAC ‘Boehm_construction X Ms p’ + >> MATCH_MP_TAC subtree_equiv_lemma_explicit' >> art [] +QED + (* Definition 10.3.10 (iii) and (iv) [1, p.251] - NOTE: The purpose of X is to make sure all terms in Ms share the same excluded + NOTE: The purpose of X is to make sure all terms in Ms share the same exclude set (and thus perhaps also the same initial binding list). *) Definition agree_upto_def : agree_upto X Ms p r <=> - !M N q. MEM M Ms /\ MEM N Ms /\ q <<= p ==> subtree_equiv X M N q r + !M N q. MEM M Ms /\ MEM N Ms /\ q <<= p ==> subtree_equiv X M N q r End -(* NOTE: subterm_equiv_lemma and this theorem together implies the original +(* NOTE: subtree_equiv_lemma and this theorem together implies the original agree_upto_lemma (see below). *) -Theorem subterm_equiv_imp_agree_upto : +Theorem subtree_equiv_imp_agree_upto : !X Ms p r pi. - (!M N q. - MEM M Ms /\ MEM N Ms /\ q <<= p /\ subtree_equiv X M N q r ==> - subtree_equiv X (apply pi M) (apply pi N) q r) /\ + (!M N q. MEM M Ms /\ MEM N Ms /\ q <<= p /\ subtree_equiv X M N q r ==> + subtree_equiv X (apply pi M) (apply pi N) q r) /\ agree_upto X Ms p r ==> agree_upto X (MAP (apply pi) Ms) p r Proof RW_TAC std_ss [agree_upto_def, MEM_MAP] @@ -55,26 +316,102 @@ Proof >> FIRST_X_ASSUM MATCH_MP_TAC >> art [] QED +(* Lemma 10.3.11 [1. p.251] *) Theorem agree_upto_lemma : !X Ms p r. FINITE X /\ p <> [] /\ 0 < r /\ Ms <> [] /\ BIGUNION (IMAGE FV (set Ms)) SUBSET X UNION RANK r /\ EVERY (\M. subterm X M p r <> NONE) Ms ==> - ?pi. Boehm_transform pi /\ EVERY is_ready' (MAP (apply pi) Ms) /\ - (agree_upto X Ms p r ==> agree_upto X (MAP (apply pi) Ms) p r) /\ - !M N. MEM M Ms /\ MEM N Ms /\ - subtree_equiv X (apply pi M) (apply pi N) p r ==> - subtree_equiv' X M N p r + ?pi. Boehm_transform pi /\ EVERY is_ready (MAP (apply pi) Ms) /\ + (agree_upto X Ms p r ==> agree_upto X (MAP (apply pi) Ms) p r) /\ + !M N. MEM M Ms /\ MEM N Ms /\ + subtree_equiv X (apply pi M) (apply pi N) p r ==> + subtree_equiv X M N p r Proof rpt GEN_TAC - >> DISCH_THEN (MP_TAC o (MATCH_MP subterm_equiv_lemma)) + >> DISCH_THEN (MP_TAC o (MATCH_MP subtree_equiv_lemma)) >> rpt STRIP_TAC - >> Q.EXISTS_TAC ‘pi’ >> rw [] - >> MATCH_MP_TAC subterm_equiv_imp_agree_upto >> art [] + >> Q.EXISTS_TAC ‘pi’ + >> RW_TAC std_ss [] + >- (MATCH_MP_TAC subtree_equiv_imp_agree_upto >> rw [] \\ + METIS_TAC []) + >> Q.PAT_X_ASSUM ‘!M N q. MEM M Ms /\ MEM N Ms /\ q <<= p ==> _’ + (MP_TAC o Q.SPECL [‘M’, ‘N’, ‘p’]) + >> simp [] +QED + +(* Definition 10.3.10 (ii) [1, p.251] *) +Definition is_faithful_def : + is_faithful p X Ms pi r <=> + (!M. MEM M Ms ==> + (solvable (apply pi M) <=> p IN ltree_paths (BT' X M r))) /\ + !M N. MEM M Ms /\ MEM N Ms ==> + (subtree_equiv X M N p r <=> equivalent (apply pi M) (apply pi N)) +End + +Overload is_faithful' = “is_faithful []” + +Theorem is_faithful' : + !X Ms pi r. Boehm_transform pi ==> + (is_faithful' X Ms pi r <=> + EVERY solvable Ms /\ + EVERY solvable (MAP (apply pi) Ms) /\ + !M N. MEM M Ms /\ MEM N Ms ==> + (subtree_equiv X M N [] r <=> + equivalent (apply pi M) (apply pi N))) +Proof + rw [ltree_paths_def, is_faithful_def, EVERY_MEM] + >> reverse EQ_TAC + >- (rw [MEM_MAP] \\ + FIRST_X_ASSUM MATCH_MP_TAC \\ + Q.EXISTS_TAC ‘M’ >> art []) + >> simp [MEM_MAP] + >> STRIP_TAC + >> ONCE_REWRITE_TAC [CONJ_COMM] + >> CONJ_ASM1_TAC + >- (rw [] >> FIRST_X_ASSUM MATCH_MP_TAC >> art []) + >> Q.X_GEN_TAC ‘M’ + >> DISCH_TAC + >> Suff ‘solvable (apply pi M)’ + >- METIS_TAC [Boehm_transform_of_unsolvables] + >> FIRST_X_ASSUM MATCH_MP_TAC + >> Q.EXISTS_TAC ‘M’ >> art [] +QED + +Theorem is_faithful_two : + !p X M N pi r. + is_faithful p X [M; N] pi r <=> + (solvable (apply pi M) <=> p IN ltree_paths (BT' X M r)) /\ + (solvable (apply pi N) <=> p IN ltree_paths (BT' X N r)) /\ + (subtree_equiv X M N p r <=> equivalent (apply pi M) (apply pi N)) +Proof + rw [is_faithful_def] + >> EQ_TAC >> rw [] (* 6 subgoals here *) + >> rw [] (* only one subgoal is left *) + >> rw [Once subtree_equiv_comm, Once equivalent_comm] +QED + +Theorem is_faithful_two' : + !p X M N pi r. + p IN ltree_paths (BT' X M r) /\ + p IN ltree_paths (BT' X N r) ==> + (is_faithful p X [M; N] pi r <=> + solvable (apply pi M) /\ + solvable (apply pi N) /\ + (subtree_equiv X M N p r <=> equivalent (apply pi M) (apply pi N))) +Proof + rw [is_faithful_two] +QED + +Theorem is_faithful_swap : + !p X M N pi r. is_faithful p X [M; N] pi r <=> is_faithful p X [N; M] pi r +Proof + rw [is_faithful_def] + >> METIS_TAC [] QED (*---------------------------------------------------------------------------* - * Separability of terms + * Separability of lambda terms *---------------------------------------------------------------------------*) Theorem separability_lemma0_case2[local] : @@ -555,9 +892,9 @@ Proof >> STRONG_CONJ_TAC >- rw [Abbr ‘pi’, Boehm_transform_def, EVERY_SNOC, EVERY_MAP] >> DISCH_TAC - (* applying Boehm_apply_unsolvable *) + (* applying Boehm_transform_of_unsolvables *) >> reverse CONJ_TAC - >- (MATCH_MP_TAC Boehm_apply_unsolvable >> art []) + >- (MATCH_MP_TAC Boehm_transform_of_unsolvables >> art []) (* stage work *) >> MATCH_MP_TAC lameq_TRANS >> Q.EXISTS_TAC ‘apply pi M0’ diff --git a/examples/lambda/basics/basic_swapScript.sml b/examples/lambda/basics/basic_swapScript.sml index 7d9546a6aa..c9d0ea8b32 100644 --- a/examples/lambda/basics/basic_swapScript.sml +++ b/examples/lambda/basics/basic_swapScript.sml @@ -222,6 +222,16 @@ Proof rw [RNEWS, alloc_prefix] QED +Theorem TAKE_RNEWS : + !r m n s. FINITE s /\ m <= n ==> TAKE m (RNEWS r n s) = RNEWS r m s +Proof + rw [RNEWS, alloc_def] + >> qabbrev_tac ‘z = string_width s’ + >> simp [Once LIST_EQ_REWRITE, listRangeLHI_def] + >> Q.X_GEN_TAC ‘i’ >> DISCH_TAC + >> simp [EL_TAKE, EL_MAP] +QED + Theorem RNEWS_set : !r n s. set (RNEWS r n s) = {v | ?j. v = n2s (r *, j) /\ diff --git a/examples/lambda/basics/termScript.sml b/examples/lambda/basics/termScript.sml index 5e0d094f8d..9f5b65af63 100644 --- a/examples/lambda/basics/termScript.sml +++ b/examples/lambda/basics/termScript.sml @@ -742,6 +742,16 @@ Definition DOM_DEF : (DOM ((x,y)::rst) = {y} UNION DOM rst) End +Theorem DOM_SNOC : + !x y rst. DOM (SNOC (x,y) rst) = {y} UNION DOM rst +Proof + NTAC 2 GEN_TAC + >> Induct_on ‘rst’ >- rw [DOM_DEF] + >> simp [FORALL_PROD] + >> rw [DOM_DEF] + >> SET_TAC [] +QED + Theorem DOM_ALT_MAP_SND : !phi. DOM phi = set (MAP SND phi) Proof @@ -756,6 +766,16 @@ Definition FVS_DEF : (FVS ((t,x)::rst) = FV t UNION FVS rst) End +Theorem FVS_SNOC : + !t x rst. FVS (SNOC (t,x) rst) = FV t UNION FVS rst +Proof + NTAC 2 GEN_TAC + >> Induct_on ‘rst’ >- rw [FVS_DEF] + >> simp [FORALL_PROD] + >> rw [FVS_DEF] + >> SET_TAC [] +QED + Theorem FVS_ALT : !ss. FVS ss = BIGUNION (set (MAP (FV o FST) ss)) Proof @@ -946,6 +966,14 @@ Proof >> CCONTR_TAC >> gs [] QED +Theorem ISUB_SNOC : + !s x rst t. t ISUB SNOC (s,x) rst = [s/x] (t ISUB rst) +Proof + NTAC 2 GEN_TAC + >> Induct_on ‘rst’ >- rw [] + >> simp [FORALL_PROD] +QED + (* ---------------------------------------------------------------------- RENAMING: a special iterated substitutions like tpm ---------------------------------------------------------------------- *) diff --git a/examples/logic/propositional_logic/Holmakefile b/examples/logic/propositional_logic/Holmakefile index f1de1b25a9..b7b42737e8 100644 --- a/examples/logic/propositional_logic/Holmakefile +++ b/examples/logic/propositional_logic/Holmakefile @@ -1,2 +1,2 @@ -INCDIRS = bag string +INCDIRS = bag string finite_maps INCLUDES = $(patsubst %,$(dprot $(HOLDIR)/src/%),$(INCDIRS)) diff --git a/examples/machine-code/garbage-collectors/arm_cheney_allocScript.sml b/examples/machine-code/garbage-collectors/arm_cheney_allocScript.sml index bd95cb440a..b6a7ce5fa8 100644 --- a/examples/machine-code/garbage-collectors/arm_cheney_allocScript.sml +++ b/examples/machine-code/garbage-collectors/arm_cheney_allocScript.sml @@ -9,8 +9,6 @@ open addressTheory mc_tailrecLib tailrecTheory; open cheney_gcTheory cheney_allocTheory arm_cheney_gcTheory; -infix \\ << >> -val op \\ = op THEN; val _ = map Parse.hide ["r0","r1","r2","r3","r4","r5","r6","r7","r8","r9","r10","r11","r12","r13"]; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/garbage-collectors/arm_cheney_gcScript.sml b/examples/machine-code/garbage-collectors/arm_cheney_gcScript.sml index d019d5f15c..73dfa04a7c 100644 --- a/examples/machine-code/garbage-collectors/arm_cheney_gcScript.sml +++ b/examples/machine-code/garbage-collectors/arm_cheney_gcScript.sml @@ -13,9 +13,6 @@ val decompile_arm = decompile arm_tools; val basic_decompile_arm = basic_decompile arm_tools; - -infix \\ << >> -val op \\ = op THEN; val _ = map Parse.hide ["r0","r1","r2","r3","r4","r5","r6","r7","r8","r9","r10","r11","r12","r13"]; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/garbage-collectors/arm_improved_gcScript.sml b/examples/machine-code/garbage-collectors/arm_improved_gcScript.sml index 2ece8896a6..6466eae802 100644 --- a/examples/machine-code/garbage-collectors/arm_improved_gcScript.sml +++ b/examples/machine-code/garbage-collectors/arm_improved_gcScript.sml @@ -9,8 +9,6 @@ open set_sepTheory bitTheory; open improved_gcTheory; open compilerLib; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/machine-code/garbage-collectors/cheney_allocScript.sml b/examples/machine-code/garbage-collectors/cheney_allocScript.sml index acfd9445f7..b0cfc8247f 100644 --- a/examples/machine-code/garbage-collectors/cheney_allocScript.sml +++ b/examples/machine-code/garbage-collectors/cheney_allocScript.sml @@ -5,8 +5,6 @@ open cheney_gcTheory; val _ = new_theory "cheney_alloc"; val _ = ParseExtras.temp_loose_equality() -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/garbage-collectors/cheney_gcScript.sml b/examples/machine-code/garbage-collectors/cheney_gcScript.sml index 5733bba2b7..1c273c33d5 100644 --- a/examples/machine-code/garbage-collectors/cheney_gcScript.sml +++ b/examples/machine-code/garbage-collectors/cheney_gcScript.sml @@ -5,8 +5,7 @@ open pred_setTheory arithmeticTheory pairTheory listTheory combinTheory finite_m val _ = new_theory "cheney_gc"; val _ = ParseExtras.temp_loose_equality() -infix \\ -val op \\ = op THEN; + val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; val bool_ss = bool_ss -* ["lift_disj_eq", "list_imp_disj"] diff --git a/examples/machine-code/garbage-collectors/improved_gcScript.sml b/examples/machine-code/garbage-collectors/improved_gcScript.sml index 3782600fdf..23be04d2e6 100644 --- a/examples/machine-code/garbage-collectors/improved_gcScript.sml +++ b/examples/machine-code/garbage-collectors/improved_gcScript.sml @@ -3,8 +3,6 @@ val _ = ParseExtras.temp_loose_equality() open pred_setTheory arithmeticTheory pairTheory listTheory combinTheory; open finite_mapTheory sumTheory relationTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/machine-code/garbage-collectors/lisp_gcScript.sml b/examples/machine-code/garbage-collectors/lisp_gcScript.sml index 7dc02d849e..6634e4f57d 100644 --- a/examples/machine-code/garbage-collectors/lisp_gcScript.sml +++ b/examples/machine-code/garbage-collectors/lisp_gcScript.sml @@ -11,8 +11,6 @@ open mc_tailrecLib tailrecTheory; open cheney_gcTheory cheney_allocTheory; (* an abstract implementation is imported *) -infix \\ << >> -val op \\ = op THEN; val _ = map Parse.hide ["r0","r1","r2","r3","r4","r5","r6","r7","r8","r9","r10","r11","r12","r13"]; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/hoare-triple/tailrecScript.sml b/examples/machine-code/hoare-triple/tailrecScript.sml index 4566bea188..d6743fffd9 100644 --- a/examples/machine-code/hoare-triple/tailrecScript.sml +++ b/examples/machine-code/hoare-triple/tailrecScript.sml @@ -4,10 +4,6 @@ open pred_setTheory arithmeticTheory whileTheory sumTheory; val _ = new_theory "tailrec"; -infix \\ -val op \\ = op THEN; - - (* ---- definitions ----- *) val TAILREC_PRE_def = Define ` diff --git a/examples/machine-code/instruction-set-models/arm/prog_armLib.sml b/examples/machine-code/instruction-set-models/arm/prog_armLib.sml index 19fb8c2bf0..817bc893de 100644 --- a/examples/machine-code/instruction-set-models/arm/prog_armLib.sml +++ b/examples/machine-code/instruction-set-models/arm/prog_armLib.sml @@ -16,10 +16,6 @@ structure Parse = struct end open Parse - -infix \\ -val op \\ = op THEN; - val use_stack = ref false; fun arm_use_stack b = (use_stack := b); diff --git a/examples/machine-code/instruction-set-models/arm/prog_armScript.sml b/examples/machine-code/instruction-set-models/arm/prog_armScript.sml index 35550ea8e4..bf71fbf971 100644 --- a/examples/machine-code/instruction-set-models/arm/prog_armScript.sml +++ b/examples/machine-code/instruction-set-models/arm/prog_armScript.sml @@ -9,9 +9,6 @@ open armTheory arm_coretypesTheory arm_stepTheory armLib; val _ = new_theory "prog_arm"; -infix \\ -val op \\ = op THEN; - val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/instruction-set-models/ppc/prog_ppcScript.sml b/examples/machine-code/instruction-set-models/ppc/prog_ppcScript.sml index 5bfe17f197..11b88e004a 100644 --- a/examples/machine-code/instruction-set-models/ppc/prog_ppcScript.sml +++ b/examples/machine-code/instruction-set-models/ppc/prog_ppcScript.sml @@ -8,10 +8,6 @@ open set_sepTheory progTheory ppc_Theory ppc_seq_monadTheory; val _ = new_theory "prog_ppc"; val _ = ParseExtras.temp_loose_equality() - -infix \\ -val op \\ = op THEN; - val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/instruction-set-models/x86/prog_x86Lib.sml b/examples/machine-code/instruction-set-models/x86/prog_x86Lib.sml index fd59d6e81b..e2a7b7306d 100644 --- a/examples/machine-code/instruction-set-models/x86/prog_x86Lib.sml +++ b/examples/machine-code/instruction-set-models/x86/prog_x86Lib.sml @@ -7,11 +7,6 @@ open set_sepTheory x86_Theory x86_Lib helperLib; open x86_seq_monadTheory x86_coretypesTheory x86_astTheory x86_icacheTheory; open prog_x86Theory; -infix \\ -val op \\ = op THEN; - - - val x86_status = xS_HIDE; val x86_pc = ``xPC``; val x86_exec_flag = ref false; diff --git a/examples/machine-code/instruction-set-models/x86/prog_x86Script.sml b/examples/machine-code/instruction-set-models/x86/prog_x86Script.sml index a186e03049..5d0799c8d5 100644 --- a/examples/machine-code/instruction-set-models/x86/prog_x86Script.sml +++ b/examples/machine-code/instruction-set-models/x86/prog_x86Script.sml @@ -8,10 +8,6 @@ open set_sepTheory progTheory x86_Theory x86_seq_monadTheory x86_icacheTheory; val _ = new_theory "prog_x86"; val _ = ParseExtras.temp_loose_equality() - -infix \\ -val op \\ = op THEN; - val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/instruction-set-models/x86_64/prog_x64Lib.sml b/examples/machine-code/instruction-set-models/x86_64/prog_x64Lib.sml index 922c6b555b..c60d10663f 100644 --- a/examples/machine-code/instruction-set-models/x86_64/prog_x64Lib.sml +++ b/examples/machine-code/instruction-set-models/x86_64/prog_x64Lib.sml @@ -17,9 +17,6 @@ end open Parse -infix \\ -val op \\ = op THEN; - val x64_status = zS_HIDE; val x64_pc = ``zPC``; val x64_exec_flag = ref true; diff --git a/examples/machine-code/instruction-set-models/x86_64/prog_x64Script.sml b/examples/machine-code/instruction-set-models/x86_64/prog_x64Script.sml index 748bf88654..269dc6bcf9 100644 --- a/examples/machine-code/instruction-set-models/x86_64/prog_x64Script.sml +++ b/examples/machine-code/instruction-set-models/x86_64/prog_x64Script.sml @@ -10,9 +10,6 @@ open temporalTheory; val _ = new_theory "prog_x64"; val _ = ParseExtras.temp_loose_equality() -infix \\ -val op \\ = op THEN; - val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/instruction-set-models/x86_64/prog_x64_extraScript.sml b/examples/machine-code/instruction-set-models/x86_64/prog_x64_extraScript.sml index f7a0cf9d9b..7c4d71b0bd 100644 --- a/examples/machine-code/instruction-set-models/x86_64/prog_x64_extraScript.sml +++ b/examples/machine-code/instruction-set-models/x86_64/prog_x64_extraScript.sml @@ -10,8 +10,6 @@ open helperLib progTheory set_sepTheory addressTheory temporalTheory; open wordsTheory wordsLib listTheory arithmeticTheory; open whileTheory pairTheory relationTheory combinTheory optionTheory; -infix \\ val op \\ = op THEN; - (* generic code gen infrastructure *) val zCODE_HEAP_RAX_def = Define ` diff --git a/examples/machine-code/just-in-time/jit_codegenScript.sml b/examples/machine-code/just-in-time/jit_codegenScript.sml index cefbda1c45..2c7ca2fee3 100644 --- a/examples/machine-code/just-in-time/jit_codegenScript.sml +++ b/examples/machine-code/just-in-time/jit_codegenScript.sml @@ -13,10 +13,6 @@ open jit_inputTheory jit_opsTheory; open compilerLib; open export_codeLib; -infix \\ -val op \\ = op THEN; - - val _ = new_theory "jit_codegen"; diff --git a/examples/machine-code/just-in-time/jit_incrementalScript.sml b/examples/machine-code/just-in-time/jit_incrementalScript.sml index 74d59ea647..e5c3fab618 100644 --- a/examples/machine-code/just-in-time/jit_incrementalScript.sml +++ b/examples/machine-code/just-in-time/jit_incrementalScript.sml @@ -12,10 +12,6 @@ open jit_codegenTheory; open compilerLib; open export_codeLib; -infix \\ -val op \\ = op THEN; - - val _ = new_theory "jit_incremental"; val _ = ParseExtras.temp_loose_equality() val bool_ss = bool_ss -* ["lift_disj_eq", "list_imp_disj"] diff --git a/examples/machine-code/just-in-time/jit_inputScript.sml b/examples/machine-code/just-in-time/jit_inputScript.sml index fdb46dad1e..00a75a15cf 100644 --- a/examples/machine-code/just-in-time/jit_inputScript.sml +++ b/examples/machine-code/just-in-time/jit_inputScript.sml @@ -6,8 +6,6 @@ open wordsLib; val _ = new_theory "jit_input"; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/just-in-time/jit_opsScript.sml b/examples/machine-code/just-in-time/jit_opsScript.sml index 11938f708f..a3269258c1 100644 --- a/examples/machine-code/just-in-time/jit_opsScript.sml +++ b/examples/machine-code/just-in-time/jit_opsScript.sml @@ -6,10 +6,6 @@ open wordsLib x86_encodeLib helperLib; open jit_inputTheory; -infix \\ -val op \\ = op THEN; - - val _ = new_theory "jit_ops"; val _ = prog_x86Lib.set_x86_code_write_perm_flag true; diff --git a/examples/machine-code/lisp/lisp_equalScript.sml b/examples/machine-code/lisp/lisp_equalScript.sml index 5b5bb1236d..e5174858c4 100644 --- a/examples/machine-code/lisp/lisp_equalScript.sml +++ b/examples/machine-code/lisp/lisp_equalScript.sml @@ -9,8 +9,6 @@ open mc_tailrecLib tailrecTheory cheney_gcTheory cheney_allocTheory; open lisp_gcTheory lisp_typeTheory lisp_invTheory; -infix \\ -val op \\ = op THEN; val _ = map Parse.hide ["r0","r1","r2","r3","r4","r5","r6","r7","r8","r9","r10","r11","r12","r13"]; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/lisp/lisp_invScript.sml b/examples/machine-code/lisp/lisp_invScript.sml index 0752cc18b9..09792d49f0 100644 --- a/examples/machine-code/lisp/lisp_invScript.sml +++ b/examples/machine-code/lisp/lisp_invScript.sml @@ -6,8 +6,6 @@ open combinTheory finite_mapTheory stringTheory addressTheory; open lisp_gcTheory cheney_gcTheory cheney_allocTheory; open lisp_typeTheory divideTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/lisp/lisp_opsScript.sml b/examples/machine-code/lisp/lisp_opsScript.sml index 9c44df3f69..627f1a3a2a 100644 --- a/examples/machine-code/lisp/lisp_opsScript.sml +++ b/examples/machine-code/lisp/lisp_opsScript.sml @@ -12,8 +12,6 @@ val decompile_arm = decompile prog_armLib.arm_tools; val decompile_ppc = decompile prog_ppcLib.ppc_tools; val decompile_x86 = decompile prog_x86Lib.x86_tools; -infix \\ -val op \\ = op THEN; val _ = map Parse.hide ["r0","r1","r2","r3","r4","r5","r6","r7","r8","r9","r10","r11","r12","r13"]; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/machine-code/lisp/lisp_parseScript.sml b/examples/machine-code/lisp/lisp_parseScript.sml index ca1c851125..0cf03f1c7e 100644 --- a/examples/machine-code/lisp/lisp_parseScript.sml +++ b/examples/machine-code/lisp/lisp_parseScript.sml @@ -8,8 +8,6 @@ open compilerLib; open lisp_gcTheory lisp_typeTheory lisp_invTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; val bool_ss = bool_ss -* ["lift_disj_eq", "lift_imp_disj"] diff --git a/examples/machine-code/lisp/lisp_printScript.sml b/examples/machine-code/lisp/lisp_printScript.sml index f46341443d..b31139137a 100644 --- a/examples/machine-code/lisp/lisp_printScript.sml +++ b/examples/machine-code/lisp/lisp_printScript.sml @@ -18,8 +18,6 @@ val decompile_ppc = decompile prog_ppcLib.ppc_tools; val decompile_x86 = decompile prog_x86Lib.x86_tools; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; val bool_ss = bool_ss -* ["lift_disj_eq", "lift_imp_disj"] diff --git a/examples/machine-code/lisp/lisp_typeScript.sml b/examples/machine-code/lisp/lisp_typeScript.sml index d6d9e59fdd..b6ef23748a 100644 --- a/examples/machine-code/lisp/lisp_typeScript.sml +++ b/examples/machine-code/lisp/lisp_typeScript.sml @@ -4,8 +4,6 @@ open wordsTheory arithmeticTheory wordsLib listTheory pred_setTheory pairTheory; open combinTheory finite_mapTheory addressTheory stringTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/miller/formalize/orderScript.sml b/examples/miller/formalize/orderScript.sml index 08ae90b251..db7f65f29a 100644 --- a/examples/miller/formalize/orderScript.sml +++ b/examples/miller/formalize/orderScript.sml @@ -25,20 +25,11 @@ open bossLib listTheory subtypeTools res_quanTools res_quanTheory pred_setTheory extra_pred_setTheory listContext relationTheory ho_proverTools subtypeTheory hurdUtils; -infixr 0 ++ << || THENC ORELSEC ORELSER ##; -infix 1 >>; - -val op!! = op REPEAT; -val op++ = op THEN; -val op<< = op THENL; -val op|| = op ORELSE; -val op>> = op THEN1; - (* ------------------------------------------------------------------------- *) (* Tools. *) (* ------------------------------------------------------------------------- *) -val S_TAC = !! (POP_ASSUM MP_TAC) ++ !! RESQ_STRIP_TAC; +val S_TAC = rpt (POP_ASSUM MP_TAC) >> rpt RESQ_STRIP_TAC; val (R_TAC, AR_TAC, R_TAC', AR_TAC') = SIMPLIFY_TACS list_c; @@ -75,40 +66,40 @@ val INSERT_SORTED = store_thm ("INSERT_SORTED", ``!f a l. totalorder f ==> (sorted f (insert f a l) = sorted f l)``, S_TAC - ++ AR_TAC [totalorder_def, partialorder_def, preorder_def, + >> AR_TAC [totalorder_def, partialorder_def, preorder_def, reflexive_def, transitive_def, total_def] - ++ Cases_on `sorted f l` << + >> Cases_on `sorted f l` THENL [R_TAC [] - ++ Cases_on `l` >> R_TAC [sorted_def, insert_def] - ++ R_TAC [insert_def] - ++ POP_ASSUM MP_TAC - ++ Q.SPEC_TAC (`h`, `x`) - ++ Induct_on `t` - >> (S_TAC - ++ Cases_on `f a x` - ++ AR_TAC [sorted_def, insert_def] - ++ ho_PROVE_TAC []) - ++ S_TAC - ++ Cases_on `f a x` >> AR_TAC [sorted_def] - ++ Q.PAT_X_ASSUM `!x. P x ==> Q x` (MP_TAC o Q.SPEC `h`) - ++ Q.PAT_X_ASSUM `sorted f (x::h::t)` MP_TAC - ++ R_TAC [sorted_def, insert_def] - ++ Cases_on `f a h` - ++ R_TAC [sorted_def] - ++ ho_PROVE_TAC [], + >> Cases_on `l` >- R_TAC [sorted_def, insert_def] + >> R_TAC [insert_def] + >> POP_ASSUM MP_TAC + >> Q.SPEC_TAC (`h`, `x`) + >> Induct_on `t` + >- (S_TAC + >> Cases_on `f a x` + >> AR_TAC [sorted_def, insert_def] + >> ho_PROVE_TAC []) + >> S_TAC + >> Cases_on `f a x` >- AR_TAC [sorted_def] + >> Q.PAT_X_ASSUM `!x. P x ==> Q x` (MP_TAC o Q.SPEC `h`) + >> Q.PAT_X_ASSUM `sorted f (x::h::t)` MP_TAC + >> R_TAC [sorted_def, insert_def] + >> Cases_on `f a h` + >> R_TAC [sorted_def] + >> ho_PROVE_TAC [], R_TAC [] - ++ Cases_on `l` >> AR_TAC [sorted_def, insert_def] - ++ POP_ASSUM MP_TAC - ++ Q.SPEC_TAC (`h`, `x`) - ++ Induct_on `t` >> R_TAC [sorted_def] - ++ S_TAC - ++ Q.PAT_X_ASSUM `!x. P x ==> Q x` (MP_TAC o Q.SPEC `h`) - ++ NTAC 2 (POP_ASSUM MP_TAC) - ++ R_TAC [sorted_def, insert_def] - ++ Cases_on `f a x` - ++ Cases_on `f a h` - ++ R_TAC [sorted_def] - ++ ho_PROVE_TAC []]); + >> Cases_on `l` >- AR_TAC [sorted_def, insert_def] + >> POP_ASSUM MP_TAC + >> Q.SPEC_TAC (`h`, `x`) + >> Induct_on `t` >- R_TAC [sorted_def] + >> S_TAC + >> Q.PAT_X_ASSUM `!x. P x ==> Q x` (MP_TAC o Q.SPEC `h`) + >> NTAC 2 (POP_ASSUM MP_TAC) + >> R_TAC [sorted_def, insert_def] + >> Cases_on `f a x` + >> Cases_on `f a h` + >> R_TAC [sorted_def] + >> ho_PROVE_TAC []]); (* non-interactive mode *) diff --git a/examples/miller/formalize/sequenceScript.sml b/examples/miller/formalize/sequenceScript.sml index 106318bd4d..26e2de1887 100644 --- a/examples/miller/formalize/sequenceScript.sml +++ b/examples/miller/formalize/sequenceScript.sml @@ -4,14 +4,6 @@ val _ = ParseExtras.temp_loose_equality() open bossLib arithmeticTheory extra_numTheory combinTheory hurdUtils; -infixr 0 ++ || ORELSEC; -infix 1 >>; -nonfix THEN ORELSE; - -val op++ = op THEN; -val op|| = op ORELSE; -val op>> = op THEN1; - (* ------------------------------------------------------------------------- *) (* Definitions. *) (* ------------------------------------------------------------------------- *) @@ -42,27 +34,27 @@ val STL_PARTIAL = store_thm ("STL_PARTIAL", ``!f. stl f = f o SUC``, FUN_EQ_TAC - ++ RW_TAC std_ss [stl_def, o_DEF]); + >> RW_TAC std_ss [stl_def, o_DEF]); val SCONS_SURJ = store_thm ("SCONS_SURJ", ``!x. ?h t. (x = scons h t)``, STRIP_TAC - ++ EXISTS_TAC ``shd x`` - ++ EXISTS_TAC ``stl x`` - ++ FUN_EQ_TAC - ++ Cases >> RW_TAC std_ss [scons_def, shd_def] - ++ RW_TAC std_ss [scons_def, stl_def]); + >> EXISTS_TAC ``shd x`` + >> EXISTS_TAC ``stl x`` + >> FUN_EQ_TAC + >> Cases >- RW_TAC std_ss [scons_def, shd_def] + >> RW_TAC std_ss [scons_def, stl_def]); val SHD_STL_ISO = store_thm ("SHD_STL_ISO", ``!h t. ?x. (shd x = h) /\ (stl x = t)``, REPEAT STRIP_TAC - ++ Q.EXISTS_TAC `\x. num_CASE x h t` - ++ RW_TAC arith_ss [shd_def] - ++ MATCH_MP_TAC EQ_EXT - ++ Cases >> RW_TAC std_ss [stl_def] - ++ RW_TAC std_ss [stl_def]); + >> Q.EXISTS_TAC `\x. num_CASE x h t` + >> RW_TAC arith_ss [shd_def] + >> MATCH_MP_TAC EQ_EXT + >> Cases >- RW_TAC std_ss [stl_def] + >> RW_TAC std_ss [stl_def]); val SHD_SCONS = store_thm ("SHD_SCONS", @@ -72,8 +64,8 @@ val SHD_SCONS = store_thm val STL_SCONS = store_thm ("STL_SCONS", ``!h t. stl (scons h t) = t``, - Suff `!h t n. stl (scons h t) n = t n` >> PROVE_TAC [EQ_EXT] - ++ RW_TAC arith_ss [stl_def, scons_def]); + Suff `!h t n. stl (scons h t) n = t n` >- PROVE_TAC [EQ_EXT] + >> RW_TAC arith_ss [stl_def, scons_def]); val SHD_SCONST = store_thm ("SHD_SCONST", @@ -84,63 +76,63 @@ val STL_SCONST = store_thm ("STL_SCONST", ``!b. stl (sconst b) = sconst b``, STRIP_TAC - ++ FUN_EQ_TAC - ++ RW_TAC std_ss [sconst_def, K_DEF, stl_def]); + >> FUN_EQ_TAC + >> RW_TAC std_ss [sconst_def, K_DEF, stl_def]); val SCONS_SHD_STL = store_thm ("SCONS_SHD_STL", ``!x. scons (shd x) (stl x) = x``, STRIP_TAC - ++ FUN_EQ_TAC - ++ Cases >> RW_TAC std_ss [scons_def, shd_def] - ++ RW_TAC std_ss [scons_def, stl_def]); + >> FUN_EQ_TAC + >> Cases >- RW_TAC std_ss [scons_def, shd_def] + >> RW_TAC std_ss [scons_def, stl_def]); val FST_o_SDEST = store_thm ("FST_o_SDEST", ``FST o sdest = shd``, FUN_EQ_TAC - ++ RW_TAC std_ss [sdest_def, o_THM]); + >> RW_TAC std_ss [sdest_def, o_THM]); val SND_o_SDEST = store_thm ("SND_o_SDEST", ``SND o sdest = stl``, FUN_EQ_TAC - ++ RW_TAC std_ss [sdest_def, o_THM]); + >> RW_TAC std_ss [sdest_def, o_THM]); val SEQUENCE_DEFINE = store_thm ("SEQUENCE_DEFINE", ``!phd ptl. ?g. (!(x : 'a). shd (g x) = phd x) /\ (!(x : 'a). stl (g x) = g (ptl x))``, RW_TAC std_ss [] - ++ Q.EXISTS_TAC `\x n. phd (FUNPOW ptl n x)` - ++ FUN_EQ_TAC - ++ RW_TAC std_ss [shd_def, stl_def, FUNPOW]); + >> Q.EXISTS_TAC `\x n. phd (FUNPOW ptl n x)` + >> FUN_EQ_TAC + >> RW_TAC std_ss [shd_def, stl_def, FUNPOW]); val SCONS_EQ = store_thm ("SCONS_EQ", ``!x xs y ys. (scons x xs = scons y ys) = (x = y) /\ (xs = ys)``, RW_TAC std_ss [] - ++ REVERSE EQ_TAC >> PROVE_TAC [] - ++ PROVE_TAC [SHD_SCONS, STL_SCONS]); + >> REVERSE EQ_TAC >- PROVE_TAC [] + >> PROVE_TAC [SHD_SCONS, STL_SCONS]); val STL_o_SDROP = store_thm ("STL_o_SDROP", ``!n. stl o sdrop n = sdrop (SUC n)``, - Induct >> RW_TAC bool_ss [sdrop_def, I_o_ID] - ++ RW_TAC bool_ss [sdrop_def, o_ASSOC]); + Induct >- RW_TAC bool_ss [sdrop_def, I_o_ID] + >> RW_TAC bool_ss [sdrop_def, o_ASSOC]); val SDROP_ADD = store_thm ("SDROP_ADD", ``!s x y. sdrop (x + y) s = sdrop x (sdrop y s)``, - Induct_on `y` >> RW_TAC list_ss [sdrop_def, I_THM] - ++ RW_TAC std_ss [ADD_CLAUSES, sdrop_def, o_THM]); + Induct_on `y` >- RW_TAC list_ss [sdrop_def, I_THM] + >> RW_TAC std_ss [ADD_CLAUSES, sdrop_def, o_THM]); val SDROP_EQ_MONO = store_thm ("SDROP_EQ_MONO", ``!m n x y. (sdrop m x = sdrop m y) /\ m <= n ==> (sdrop n x = sdrop n y)``, RW_TAC std_ss [LESS_EQ_EXISTS] - ++ Induct_on `p` >> RW_TAC arith_ss [] - ++ RW_TAC std_ss [ADD_CLAUSES, GSYM STL_o_SDROP, o_THM]); + >> Induct_on `p` >- RW_TAC arith_ss [] + >> RW_TAC std_ss [ADD_CLAUSES, GSYM STL_o_SDROP, o_THM]); val EVENTUALLY_REFL = store_thm ("EVENTUALLY_REFL", @@ -151,34 +143,34 @@ val EVENTUALLY_SYM = store_thm ("EVENTUALLY_SYM", ``!x y. eventually x y = eventually y x``, RW_TAC std_ss [eventually_def] - ++ PROVE_TAC []); + >> PROVE_TAC []); val EVENTUALLY_TRANS = store_thm ("EVENTUALLY_TRANS", ``!x y z. eventually x y /\ eventually y z ==> eventually x z``, RW_TAC std_ss [eventually_def] - ++ Q.EXISTS_TAC `MAX n n'` - ++ PROVE_TAC [X_LE_MAX, LESS_EQ_REFL, SDROP_EQ_MONO]); + >> Q.EXISTS_TAC `MAX n n'` + >> PROVE_TAC [X_LE_MAX, LESS_EQ_REFL, SDROP_EQ_MONO]); val SEQUENCE_DEFINE_ALT = store_thm ("SEQUENCE_DEFINE_ALT", ``!phd ptl. ?g:'a->num->'b. !x. g x = scons (phd x) (g (ptl x))``, RW_TAC std_ss [] - ++ MP_TAC (Q.SPECL [`phd`, `ptl`] SEQUENCE_DEFINE) - ++ RW_TAC std_ss [] - ++ Q.EXISTS_TAC `g` - ++ RW_TAC std_ss [] - ++ MP_TAC (Q.ISPEC `(g:'a->num->'b) x` SCONS_SURJ) - ++ RW_TAC std_ss [] - ++ Q.PAT_X_ASSUM `!x. P x` (MP_TAC o Q.SPEC `x`) - ++ Q.PAT_X_ASSUM `!x. P x` (MP_TAC o Q.SPEC `x`) - ++ RW_TAC std_ss [SCONS_EQ, SHD_SCONS, STL_SCONS]); + >> MP_TAC (Q.SPECL [`phd`, `ptl`] SEQUENCE_DEFINE) + >> RW_TAC std_ss [] + >> Q.EXISTS_TAC `g` + >> RW_TAC std_ss [] + >> MP_TAC (Q.ISPEC `(g:'a->num->'b) x` SCONS_SURJ) + >> RW_TAC std_ss [] + >> Q.PAT_X_ASSUM `!x. P x` (MP_TAC o Q.SPEC `x`) + >> Q.PAT_X_ASSUM `!x. P x` (MP_TAC o Q.SPEC `x`) + >> RW_TAC std_ss [SCONS_EQ, SHD_SCONS, STL_SCONS]); local val th = prove (``?s. !h t x. s h t x = scons (h x) (s h t (t x))``, MP_TAC SEQUENCE_DEFINE_ALT - ++ RW_TAC std_ss [SKOLEM_THM]) + >> RW_TAC std_ss [SKOLEM_THM]) in val siter_def = new_specification ("siter_def", ["siter"], th); end; @@ -187,13 +179,13 @@ val SHD_SITER = store_thm ("SHD_SITER", ``!h t x. shd (siter h t x) = h x``, ONCE_REWRITE_TAC [siter_def] - ++ RW_TAC std_ss [SHD_SCONS]); + >> RW_TAC std_ss [SHD_SCONS]); val STL_SITER = store_thm ("STL_SITER", ``!h t x. stl (siter h t x) = siter h t (t x)``, RW_TAC std_ss [] - ++ CONV_TAC (LAND_CONV (ONCE_REWRITE_CONV [siter_def])) - ++ RW_TAC std_ss [STL_SCONS]); + >> CONV_TAC (LAND_CONV (ONCE_REWRITE_CONV [siter_def])) + >> RW_TAC std_ss [STL_SCONS]); val _ = export_theory (); diff --git a/examples/pgcl/src/relScript.sml b/examples/pgcl/src/relScript.sml index 17df51da0e..63422764be 100644 --- a/examples/pgcl/src/relScript.sml +++ b/examples/pgcl/src/relScript.sml @@ -34,13 +34,6 @@ val _ = new_theory "rel"; (* Helpful proof tools *) (* ------------------------------------------------------------------------- *) -infixr 0 ++ << || THENC ORELSEC ORELSER ##; -infix 1 >>; - -val op ++ = op THEN; -val op << = op THENL; -val op >> = op THEN1; -val op || = op ORELSE; val Know = Q_TAC KNOW_TAC; val Suff = Q_TAC SUFF_TAC; val REVERSE = Tactical.REVERSE; diff --git a/examples/simple_complexity/lib/files.txt b/examples/simple_complexity/lib/files.txt deleted file mode 100644 index 6cf141aab5..0000000000 --- a/examples/simple_complexity/lib/files.txt +++ /dev/null @@ -1,15 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Simple_Complexity Common Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2018 *) -(* ------------------------------------------------------------------------- *) - -0 bitsize -- bit size of number representations. -* divides -* gcd -* logroot - -1 complexity -- big-O notation for complexity class. -* 0 bitsize - diff --git a/examples/simple_complexity/loop/files.txt b/examples/simple_complexity/loop/files.txt deleted file mode 100644 index 4f5d68a6c8..0000000000 --- a/examples/simple_complexity/loop/files.txt +++ /dev/null @@ -1,44 +0,0 @@ -(* ------------------------------------------------------------------------- *) -(* Hierarchy of Simple_Complexity Loop Library *) -(* *) -(* Author: Joseph Chan *) -(* Date: December, 2018 *) -(* ------------------------------------------------------------------------- *) - -0 loop -- general theory of loop recurrence, with body and exit. -* list -* rich_list -* listRange - -1 loopIncrease -- recurrence theory of increasing loops. -* list -* rich_list -* listRange -* 0 loop - -1 loopDecrease -- recurrence theory of decreasing loops. -* list -* rich_list -* listRange -* 0 loop - -1 loopDivide -- recurrence theory of dividing loops. -* list -* rich_list -* listRange -* 0 loop - -1 loopMultiply -- recurrence theory of multiplying loops. -* list -* rich_list -* listRange -* logroot -* logPower -* 0 loop - -1 loopList -- recurrence theory of list reduction loops. -* list -* rich_list -* listRange -* sublist -* 0 loop diff --git a/examples/theorem-prover/lisp-runtime/bytecode/lisp_alt_semanticsScript.sml b/examples/theorem-prover/lisp-runtime/bytecode/lisp_alt_semanticsScript.sml index 328414d85e..7a7776742d 100644 --- a/examples/theorem-prover/lisp-runtime/bytecode/lisp_alt_semanticsScript.sml +++ b/examples/theorem-prover/lisp-runtime/bytecode/lisp_alt_semanticsScript.sml @@ -5,8 +5,6 @@ open optionTheory arithmeticTheory relationTheory; open lisp_sexpTheory lisp_parseTheory lisp_semanticsTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/bytecode/lisp_bytecodeScript.sml b/examples/theorem-prover/lisp-runtime/bytecode/lisp_bytecodeScript.sml index 4666188090..369a2bd8c9 100644 --- a/examples/theorem-prover/lisp-runtime/bytecode/lisp_bytecodeScript.sml +++ b/examples/theorem-prover/lisp-runtime/bytecode/lisp_bytecodeScript.sml @@ -2,8 +2,6 @@ open HolKernel boolLib bossLib Parse; val _ = new_theory "lisp_bytecode"; open stringTheory arithmeticTheory lisp_sexpTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/bytecode/lisp_compilerScript.sml b/examples/theorem-prover/lisp-runtime/bytecode/lisp_compilerScript.sml index f9d322d8d3..515323fc3b 100644 --- a/examples/theorem-prover/lisp-runtime/bytecode/lisp_compilerScript.sml +++ b/examples/theorem-prover/lisp-runtime/bytecode/lisp_compilerScript.sml @@ -7,8 +7,6 @@ open optionTheory arithmeticTheory relationTheory pairTheory; open lisp_bytecodeTheory; open lisp_sexpTheory lisp_semanticsTheory lisp_alt_semanticsTheory -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/extract/lisp_extractScript.sml b/examples/theorem-prover/lisp-runtime/extract/lisp_extractScript.sml index ff01e965bc..abc68d219a 100644 --- a/examples/theorem-prover/lisp-runtime/extract/lisp_extractScript.sml +++ b/examples/theorem-prover/lisp-runtime/extract/lisp_extractScript.sml @@ -6,8 +6,6 @@ open optionTheory arithmeticTheory relationTheory; open lisp_sexpTheory lisp_semanticsTheory lisp_parseTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/extract/lisp_synthesis_demoScript.sml b/examples/theorem-prover/lisp-runtime/extract/lisp_synthesis_demoScript.sml index 6c8798de17..1ec96d20b9 100644 --- a/examples/theorem-prover/lisp-runtime/extract/lisp_synthesis_demoScript.sml +++ b/examples/theorem-prover/lisp-runtime/extract/lisp_synthesis_demoScript.sml @@ -3,9 +3,6 @@ open HolKernel Parse boolLib bossLib; val _ = new_theory "lisp_synthesis_demo"; open arithmeticTheory listTheory pairTheory lisp_sexpTheory lisp_synthesisLib; -infix \\ val op \\ = op THEN; - - (* we start by proving a lemma which helps with termination proofs *) val term_lemma = prove( diff --git a/examples/theorem-prover/lisp-runtime/garbage-collector/lisp_consScript.sml b/examples/theorem-prover/lisp-runtime/garbage-collector/lisp_consScript.sml index 519b949d33..5a141d65be 100644 --- a/examples/theorem-prover/lisp-runtime/garbage-collector/lisp_consScript.sml +++ b/examples/theorem-prover/lisp-runtime/garbage-collector/lisp_consScript.sml @@ -9,8 +9,6 @@ open set_sepTheory bitTheory fcpTheory; open stop_and_copyTheory; open codegenLib decompilerLib prog_x64Lib prog_x64Theory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/garbage-collector/stop_and_copyScript.sml b/examples/theorem-prover/lisp-runtime/garbage-collector/stop_and_copyScript.sml index 263fbd9c51..86bc5523f3 100644 --- a/examples/theorem-prover/lisp-runtime/garbage-collector/stop_and_copyScript.sml +++ b/examples/theorem-prover/lisp-runtime/garbage-collector/stop_and_copyScript.sml @@ -4,8 +4,6 @@ val _ = ParseExtras.temp_loose_equality() open pred_setTheory arithmeticTheory pairTheory listTheory combinTheory; open finite_mapTheory sumTheory relationTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_bigopsScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_bigopsScript.sml index a062f30b48..14a43c3ef8 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_bigopsScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_bigopsScript.sml @@ -18,8 +18,6 @@ val _ = let (* --- *) -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_bytecode_stepScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_bytecode_stepScript.sml index 9405ba9aa0..641d1526b1 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_bytecode_stepScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_bytecode_stepScript.sml @@ -12,8 +12,6 @@ open combinTheory finite_mapTheory addressTheory helperLib sumTheory; open set_sepTheory bitTheory fcpTheory stringTheory optionTheory relationTheory; open stop_and_copyTheory lisp_consTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_codegenScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_codegenScript.sml index 055eacb67e..7090bb6ec2 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_codegenScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_codegenScript.sml @@ -14,8 +14,6 @@ open stop_and_copyTheory; open codegenLib decompilerLib prog_x64Lib prog_x64Theory progTheory; open lisp_parseTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_compiler_opScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_compiler_opScript.sml index 5f81a2fee4..256b8ea997 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_compiler_opScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_compiler_opScript.sml @@ -29,8 +29,6 @@ val _ = let open lisp_compilerTheory lisp_semanticsTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_correctnessScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_correctnessScript.sml index f404433875..54d26563aa 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_correctnessScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_correctnessScript.sml @@ -11,8 +11,6 @@ open wordsTheory arithmeticTheory wordsLib listTheory pred_setTheory pairTheory; open combinTheory finite_mapTheory addressTheory helperLib sumTheory; open set_sepTheory bitTheory fcpTheory stringTheory optionTheory relationTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_equalScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_equalScript.sml index c1e211be08..f11e0941dc 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_equalScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_equalScript.sml @@ -12,8 +12,6 @@ open set_sepTheory bitTheory fcpTheory stringTheory; open codegenLib decompilerLib prog_x64Lib; open lisp_consTheory stop_and_copyTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_invScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_invScript.sml index 64c9312d5a..6b120274a6 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_invScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_invScript.sml @@ -9,8 +9,6 @@ open set_sepTheory bitTheory fcpTheory; open lisp_sexpTheory lisp_consTheory stop_and_copyTheory lisp_bytecodeTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_opsScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_opsScript.sml index 384d191ecb..e33eca20ee 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_opsScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_opsScript.sml @@ -14,8 +14,6 @@ open prog_x64Theory prog_x64Lib x64_encodeLib; open stop_and_copyTheory; fun allowing_rebinds f x = Feedback.trace ("Theory.allow_rebinds", 1) f x -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/lisp-runtime/implementation/lisp_symbolsScript.sml b/examples/theorem-prover/lisp-runtime/implementation/lisp_symbolsScript.sml index 60c1963688..beb9a2a331 100644 --- a/examples/theorem-prover/lisp-runtime/implementation/lisp_symbolsScript.sml +++ b/examples/theorem-prover/lisp-runtime/implementation/lisp_symbolsScript.sml @@ -14,8 +14,6 @@ open stop_and_copyTheory; open codegenLib decompilerLib prog_x64Lib prog_x64Theory progTheory; open lisp_parseTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; fun SUBGOAL q = REVERSE (sg q) diff --git a/examples/theorem-prover/lisp-runtime/parse/lisp_parseScript.sml b/examples/theorem-prover/lisp-runtime/parse/lisp_parseScript.sml index 30f9b83058..fe1212128a 100644 --- a/examples/theorem-prover/lisp-runtime/parse/lisp_parseScript.sml +++ b/examples/theorem-prover/lisp-runtime/parse/lisp_parseScript.sml @@ -15,8 +15,6 @@ val std_ss = std_ss -* ["lift_disj_eq", "lift_imp_disj"] val bool_ss = bool_ss -* ["lift_disj_eq", "lift_imp_disj"] val _ = temp_delsimps ["lift_disj_eq", "lift_imp_disj"] -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/milawa-prover/milawa_coreScript.sml b/examples/theorem-prover/milawa-prover/milawa_coreScript.sml index 6b2629650a..12f120ec5d 100644 --- a/examples/theorem-prover/milawa-prover/milawa_coreScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_coreScript.sml @@ -5,8 +5,6 @@ open optionTheory arithmeticTheory relationTheory; open lisp_sexpTheory lisp_parseTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/milawa-prover/milawa_defsScript.sml b/examples/theorem-prover/milawa-prover/milawa_defsScript.sml index 98a8f27da7..ab583dcf18 100644 --- a/examples/theorem-prover/milawa-prover/milawa_defsScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_defsScript.sml @@ -10,8 +10,6 @@ open lisp_sexpTheory lisp_semanticsTheory lisp_parseTheory; open milawa_initTheory; val _ = max_print_depth := 20; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/milawa-prover/milawa_execScript.sml b/examples/theorem-prover/milawa-prover/milawa_execScript.sml index b9bd2c935b..e3f4fa24b3 100644 --- a/examples/theorem-prover/milawa-prover/milawa_execScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_execScript.sml @@ -6,10 +6,6 @@ open optionTheory arithmeticTheory relationTheory combinTheory; open lisp_sexpTheory lisp_semanticsTheory; open milawa_logicTheory milawa_defsTheory; -infix \\ -val op \\ = op THEN; - - (* We define an intermediate language MR_ev -- a language which is very much like the runtime's specification R_ev. The difference is that MR_ev is deterministic, functions that are simply just Error diff --git a/examples/theorem-prover/milawa-prover/milawa_initScript.sml b/examples/theorem-prover/milawa-prover/milawa_initScript.sml index 7f255b60ef..f0fc6a1e53 100644 --- a/examples/theorem-prover/milawa-prover/milawa_initScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_initScript.sml @@ -11,8 +11,6 @@ open lisp_sexpTheory lisp_semanticsTheory lisp_parseTheory; open milawa_coreTheory; val _ = max_print_depth := 20; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE; val RW1 = ONCE_REWRITE_RULE; diff --git a/examples/theorem-prover/milawa-prover/milawa_logicScript.sml b/examples/theorem-prover/milawa-prover/milawa_logicScript.sml index 0b2d4a1ae4..06556f3dd4 100644 --- a/examples/theorem-prover/milawa-prover/milawa_logicScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_logicScript.sml @@ -16,9 +16,6 @@ open pred_setTheory combinTheory finite_mapTheory stringTheory; -infix \\ -val op \\ = op THEN; - val std_ss = std_ss -* ["lift_disj_eq", "lift_imp_disj"] val bool_ss = bool_ss -* ["lift_disj_eq", "lift_imp_disj"] val _ = temp_delsimps ["lift_disj_eq", "lift_imp_disj"] diff --git a/examples/theorem-prover/milawa-prover/milawa_ordinalScript.sml b/examples/theorem-prover/milawa-prover/milawa_ordinalScript.sml index 633c5da84d..d1bf770cc0 100644 --- a/examples/theorem-prover/milawa-prover/milawa_ordinalScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_ordinalScript.sml @@ -3,8 +3,6 @@ open HolKernel Parse boolLib bossLib; val _ = new_theory "milawa_ordinal"; open lisp_sexpTheory arithmeticTheory pred_setTheory ordinalNotationTheory; -infix \\ -val op \\ = op THEN; val RW = REWRITE_RULE diff --git a/examples/theorem-prover/milawa-prover/milawa_proofpScript.sml b/examples/theorem-prover/milawa-prover/milawa_proofpScript.sml index dadd8cab62..7eee80c056 100644 --- a/examples/theorem-prover/milawa-prover/milawa_proofpScript.sml +++ b/examples/theorem-prover/milawa-prover/milawa_proofpScript.sml @@ -7,8 +7,6 @@ open milawa_defsTheory milawa_logicTheory milawa_execTheory; open arithmeticTheory listTheory pred_setTheory finite_mapTheory combinTheory; open pairTheory; -infix \\ -val op \\ = op THEN; val _ = temp_delsimps ["NORMEQ_CONV"] val _ = diminish_srw_ss ["ABBREV"] val _ = set_trace "BasicProvers.var_eq_old" 1 diff --git a/examples/theorem-prover/milawa-prover/soundness-thm/milawa_soundnessScript.sml b/examples/theorem-prover/milawa-prover/soundness-thm/milawa_soundnessScript.sml index 00ead980e1..e9c3a4db28 100644 --- a/examples/theorem-prover/milawa-prover/soundness-thm/milawa_soundnessScript.sml +++ b/examples/theorem-prover/milawa-prover/soundness-thm/milawa_soundnessScript.sml @@ -6,8 +6,6 @@ open HolKernel Parse boolLib bossLib; val _ = new_theory "milawa_soundness"; open helperLib set_sepTheory progTheory; open lisp_correctnessTheory milawa_proofpTheory lisp_semanticsTheory; -infix \\ val op \\ = op THEN - val R_exec_T_11 = prove( ``!x y. R_exec x y ==> !z. SND y /\ R_exec x z ==> (y = z)``, HO_MATCH_MP_TAC R_exec_ind \\ STRIP_TAC diff --git a/help/Docfiles/Portable.remove_external_wspace.doc b/help/Docfiles/Portable.remove_external_wspace.doc new file mode 100644 index 0000000000..41c8b97f31 --- /dev/null +++ b/help/Docfiles/Portable.remove_external_wspace.doc @@ -0,0 +1,25 @@ +\DOC + +\TYPE {remove_external_wspace : string -> string} + +\SYNOPSIS +Removes trailing and leading whitespace characters from a string + +\KEYWORDS + +\DESCRIBE +A call to {remove_external_wspace s} returns a string identical to {s} +except that all leading and trailing characters for which +{Char.isSpace} is true have been removed. The implementation is (with +the Basis's {Substring} structure open): +{ + string (dropl Char.isSpace (dropr Char.isSpace (full s))) +} + +\FAILURE +Never fails. + +\SEEALSO +Portable.remove_wspace. + +\ENDDOC diff --git a/help/Docfiles/Portable.remove_wspace.doc b/help/Docfiles/Portable.remove_wspace.doc index 8dcfc1d8fa..c04aae6b69 100644 --- a/help/Docfiles/Portable.remove_wspace.doc +++ b/help/Docfiles/Portable.remove_wspace.doc @@ -18,4 +18,7 @@ removed. The implementation is \FAILURE Never fails. +\SEEALSO +Portable.remove_external_wspace. + \ENDDOC diff --git a/src/1/ThmAttribute.sig b/src/1/ThmAttribute.sig index 0569f48476..e344f8369e 100644 --- a/src/1/ThmAttribute.sig +++ b/src/1/ThmAttribute.sig @@ -1,15 +1,22 @@ signature ThmAttribute = sig - type attrfun = {name:string,attrname:string,thm:Thm.thm} -> unit + type attrfun = + {name:string,attrname:string,args:string list,thm:Thm.thm} -> unit type attrfuns = {localf : attrfun, storedf : attrfun} val register_attribute : string * attrfuns -> unit + val reserve_word : string -> unit + val is_attribute : string -> bool val all_attributes : unit -> string list val store_at_attribute : attrfun val local_attribute : attrfun - val extract_attributes : string -> string * string list + val extract_attributes : + string -> {thmname : string, + attrs : (string * string list) list, + unknown : (string * string list) list, + reserved : (string * string list) list} val toString : string * string list -> string val insert_attribute : {attr: string} -> string -> string @@ -18,8 +25,12 @@ end (* [extract_attributes thmstr] takes a string of the form - thmname[attr1,attr2,...] or of the form thmname and returns the thmname - along with the list of attributes. + thmname[attr1,attr2,...] or of the form thmname and returns the + thmname along with the list of attributes, with attributes + classified as reserved (to be handled specially), "attrs" (those + that have functions stored for them), and unknowns. The arguments + are whitespace delimited strings that appear after an "=" sign. - [toString thmname attrs] reverses the extract_attributes function above. + [toString thmname attrs] reverses the extract_attributes function + above. *) diff --git a/src/1/ThmAttribute.sml b/src/1/ThmAttribute.sml index 769d9c3c68..be1b3b5470 100644 --- a/src/1/ThmAttribute.sml +++ b/src/1/ThmAttribute.sml @@ -2,27 +2,52 @@ structure ThmAttribute :> ThmAttribute = struct local open Symtab in end - type attrfun = {name:string,attrname:string,thm:Thm.thm} -> unit + type attrfun = {name:string,attrname:string,args:string list,thm:Thm.thm} -> + unit type attrfuns = {localf: attrfun, storedf : attrfun} structure Map = Symtab val funstore = ref (Map.empty : attrfuns Map.table) + fun all_attributes () = Map.keys (!funstore) + fun is_attribute a = Map.defined (!funstore) a + + val reserved_words = ref ["induction"] + (* "induction=name" is handled by tools/Holmake/HolParser, and so is basically + invisible to all later code *) + fun reserve_word s = + if Lib.mem s (!reserved_words) then + if !Globals.interactive then + Feedback.HOL_WARNING "ThmAttribute" "reserve_word" + ("Word \"" ^ s ^ "\" already reserved") + else + raise Feedback.mk_HOL_ERR "ThmAttribute" "reserve_word" + ("Word \"" ^ s ^ "\" already reserved") + else if is_attribute s then + raise Feedback.mk_HOL_ERR "ThmAttribute" "reserve_word" + ("Word \"" ^ s ^ "\" already a standard attribute") + else + reserved_words := s :: (!reserved_words) - val reserved_attrnames = ["local", "unlisted", "nocompute", "schematic", - "notuserdef", "allow_rebind", "tailrecursive"] +(* + "unlisted", "nocompute", "schematic", + "tailrecursive" +*) fun okchar c = Char.isAlphaNum c orelse c = #"_" orelse c = #"'" - fun illegal_attrname s = Lib.mem s reserved_attrnames orelse - String.isPrefix "induction=" s orelse - s = "" orelse - not (Char.isAlpha (String.sub(s,0))) orelse - not (CharVector.all okchar s) + fun legal_attrsyntax s = + s <> "" andalso + Char.isAlpha (String.sub(s,0)) andalso + CharVector.all okchar s fun register_attribute (kv as (s, f)) = let - val _ = not (illegal_attrname s) orelse + val _ = legal_attrsyntax s orelse raise Feedback.mk_HOL_ERR "ThmAttribute" "register_attribute" ("Illegal attribute name: \""^s^"\"") + val _ = not (Lib.mem s (!reserved_words)) orelse + raise Feedback.mk_HOL_ERR "ThmAttribute" "register_attribute" + ("Name \""^s^"\" already reserved for other uses") + val oldm = !funstore val newm = case Map.lookup oldm s of @@ -37,9 +62,7 @@ struct funstore := newm end - fun all_attributes () = Map.keys (!funstore) - fun is_attribute a = Map.defined (!funstore) a - fun at_attribute nm sel (arg as {name,attrname,thm}) = + fun at_attribute nm sel (arg as {name,attrname,args,thm}) = case Map.lookup (!funstore) attrname of NONE => raise Feedback.mk_HOL_ERR "ThmAttribute" nm ("No such attribute: " ^ attrname) @@ -47,7 +70,7 @@ struct val store_at_attribute = at_attribute "store_at_attribute" #storedf val local_attribute = at_attribute "local_attribute" #localf - fun extract_attributes s = let + fun extract_attributes0 s = let open Substring val (bracketl,rest) = position "[" (full s) in @@ -56,19 +79,53 @@ struct val (names,bracketr) = position "]" (slice(rest,1,NONE)) in if size bracketr <> 1 then - raise Feedback.mk_HOL_ERR "boolLib" "resolve_storename" + raise Feedback.mk_HOL_ERR "ThmAttribute" "extract_attributes" ("Malformed theorem-binding specifier: "^s) else - (string bracketl, String.fields (fn c => c = #",") (string names)) + (string bracketl, + map Portable.remove_external_wspace + (String.fields (fn c => c = #",") (string names))) end end + fun dest_attribute_equality astr = + let + open Substring + val (eql,rest) = position "=" (full astr) + val key = Portable.remove_external_wspace (string eql) + in + if isEmpty rest then (key, []) + else + let val vs = string (slice(rest,1,NONE)) + in + (key, + map Portable.remove_external_wspace + (String.tokens Char.isSpace vs)) + end + end + + fun extract_attributes s = + let val (thmname,rawattrs) = extract_attributes0 s + fun categorise (U,R,A) attrs = + case attrs of + [] => + {thmname = thmname, attrs = List.rev A, unknown = List.rev U, + reserved = List.rev R} + | (a as (k,vs)) :: rest => + if is_attribute k then categorise(U,R,a::A) rest + else if Lib.mem k (!reserved_words) then + categorise(U,a::R,A) rest + else categorise (a::U,R,A) rest + in + categorise([],[],[]) (map dest_attribute_equality rawattrs) + end + fun toString (s, attrs) = if null attrs then s else s ^ "[" ^ String.concatWith "," attrs ^ "]" fun insert_attribute {attr} s = - let val (s0,attrs) = extract_attributes s + let val (s0,attrs) = extract_attributes0 s in toString (s0, attr::attrs) end diff --git a/src/1/ThmSetData.sml b/src/1/ThmSetData.sml index 389f947f00..940a006f0a 100644 --- a/src/1/ThmSetData.sml +++ b/src/1/ThmSetData.sml @@ -205,10 +205,16 @@ fun new_exporter {settype = name, efns = efns as {add, remove}} = let export_deltasexp data end - fun store_attrfun {attrname,name,thm} = export (toKName name,thm) - fun local_attrfun {attrname,name,thm} = - add {named_thm = ({Thy = current_theory(), Name = name},thm), - thy = current_theory()} + fun store_attrfun {attrname,name,args,thm} = + if null args then export (toKName name,thm) + else raise ERR "store_attrfun" + ("Arguments not allowed for attribute " ^ attrname) + fun local_attrfun {attrname,name,args,thm} = + if null args then + add {named_thm = ({Thy = current_theory(), Name = name},thm), + thy = current_theory()} + else raise ERR "local_attrfun" + ("Arguments not allowed for attribute " ^ attrname) in data_map := Symtab.update(name,(segdata, efns)) (!data_map); ThmAttribute.register_attribute ( @@ -252,16 +258,23 @@ fun export_with_ancestry AncestryData.fullmake { adinfo = adinfo, uptodate_delta = uptodate, sexps = sexps, globinfo = globinfo} - fun store_attrfun {attrname,name,thm} = - let val d = rADD(lift name) - in - #record_delta fullresult d; - #update_global_value fullresult (raw_apply_global d) - end - fun local_attrfun {attrname,name,thm} = + fun store_attrfun {attrname,name,thm,args} = + if null args then + let val d = rADD(lift name) + in + #record_delta fullresult d; + #update_global_value fullresult (raw_apply_global d) + end + else raise ERR "store_attrfun" + ("Arguments not allowed for attribute " ^ attrname) + fun local_attrfun {attrname,name,thm,args} = (* as this is local, the name is not going to be a valid binding, and "cooking" a raw ADD delta will just throw an exception *) - #update_global_value fullresult(apply_to_global (ADD(lift name, thm))) + if null args then + #update_global_value fullresult(apply_to_global (ADD(lift name, thm))) + else + raise ERR "local_attrfun" + ("Arguments not allowed for attribute " ^ attrname) fun efn_add {thy,named_thm} = #update_global_value fullresult(raw_apply_global(rADD (#1 named_thm))) fun efn_remove {thy,remove} = diff --git a/src/1/boolLib.sml b/src/1/boolLib.sml index 9fe9b2f139..d44f2ffd4a 100644 --- a/src/1/boolLib.sml +++ b/src/1/boolLib.sml @@ -117,17 +117,27 @@ fun prove_local loc privp (n,th) = else (); DB.store_local {private=privp,loc=loc,class=DB_dtype.Thm} n th; th) +val _ = app ThmAttribute.reserve_word ["local", "unlisted", "allow_rebind"] fun extract_localpriv (loc,priv,rebindok,acc) attrs = case attrs of [] => (loc,priv,rebindok,List.rev acc) - | "unlisted" :: rest => extract_localpriv (loc,true,rebindok,acc) rest - | "local" :: rest => extract_localpriv (true,priv,rebindok,acc) rest - | "allow_rebind" :: rest => extract_localpriv (loc,priv,true,acc) rest + | ("unlisted",_) :: rest => extract_localpriv (loc,true,rebindok,acc) rest + | ("local",_) :: rest => extract_localpriv (true,priv,rebindok,acc) rest + | ("allow_rebind",_) :: rest => extract_localpriv (loc,priv,true,acc) rest | a :: rest => extract_localpriv (loc,priv,rebindok,a::acc) rest in -fun save_thm_attrs loc (n, attrs, th) = let - val (localp,privp,rebindok,attrs) = - extract_localpriv (false,false,false,[]) attrs +fun save_thm_attrs loc (attrblock, th) = let + val {thmname=n,attrs,unknown,reserved} = attrblock + val _ = null unknown orelse + raise mk_HOL_ERR "boolLib" "save_thm_attrs" + ("Unknown attributes: " ^ + String.concatWith ", " (map #1 unknown)) + val (localp,privp,rebindok,reserved_leftover) = + extract_localpriv (false,false,false,[]) reserved + val _ = null reserved_leftover orelse + raise mk_HOL_ERR "boolLib" "save_thm_attrs" + ("Unhandled attributes: " ^ + String.concatWith ", " (map #1 reserved_leftover)) val save = if localp then prove_local loc privp else @@ -136,24 +146,21 @@ fun save_thm_attrs loc (n, attrs, th) = let else ThmAttribute.store_at_attribute val storemod = if rebindok then trace("Theory.allow_rebinds", 1) else (fn f => f) - fun do_attr a = attrf {thm = th, name = n, attrname = a} + fun do_attr (k,vs) = attrf {thm = th, name = n, attrname = k, args = vs} in storemod save(n,th) before app do_attr attrs end fun store_thm_at loc (n0,t,tac) = let - val (n, attrs) = ThmAttribute.extract_attributes n0 + val attrblock = ThmAttribute.extract_attributes n0 val th = Tactical.prove(t,tac) - handle e => (print ("Failed to prove theorem " ^ n ^ ".\n"); + handle e => (print ("Failed to prove theorem " ^ #thmname attrblock ^ ".\n"); Raise e) in - save_thm_attrs loc (n,attrs,th) + save_thm_attrs loc (attrblock,th) end val store_thm = store_thm_at DB.Unknown -fun save_thm_at loc (n0,th) = let - val (n,attrs) = ThmAttribute.extract_attributes n0 -in - save_thm_attrs loc (n,attrs,th) -end +fun save_thm_at loc (n0,th) = + save_thm_attrs loc (ThmAttribute.extract_attributes n0,th) val save_thm = save_thm_at DB.Unknown fun new_recursive_definition rcd = diff --git a/src/1/boolSyntax.sml b/src/1/boolSyntax.sml index 3c023d34ba..a89633793b 100644 --- a/src/1/boolSyntax.sml +++ b/src/1/boolSyntax.sml @@ -300,9 +300,12 @@ fun defname t = fst (dest_var head handle HOL_ERR _ => dest_const head) end -fun test_remove s [] = (false, []) - | test_remove s (t::ts) = if s = t then (true, Lib.set_diff ts [s]) - else apsnd (cons t) $ test_remove s ts +fun test_remove s kvs = + let + val (ss,rest) = Portable.partition (fn (k,_) => k = s) kvs + in + (not (null ss), rest) + end fun bogus_attr cstr cnm a = HOL_WARNING cstr cnm ("No sense in " ^ a ^ " attribute on def'n") @@ -310,27 +313,34 @@ fun bogus_attr cstr cnm a = fun remove_junk cstr cnm junkas attrs0 = let fun recurse [] = [] - | recurse (a::t) = if mem a junkas then (bogus_attr cstr cnm a; - recurse t) - else a::recurse t + | recurse ((a as (k,_))::t) = + if mem k junkas then (bogus_attr cstr cnm k; recurse t) + else a::recurse t in recurse attrs0 end +val _ = List.app ThmAttribute.reserve_word ["notuserdef"] fun new_thm_with_attributes {call_str, call_f} genth (s, arg) = let open ThmAttribute - val (s0,attrs) = ThmAttribute.extract_attributes s - val (notuserdefp, attrs) = test_remove "notuserdef" attrs - val attrs = remove_junk call_str call_f - ["local", "schematic", "nocompute", "unlisted"] - attrs + val {thmname=s0,reserved=R,unknown=U,attrs=attrs} = + ThmAttribute.extract_attributes s + val (notuserdefp, R) = test_remove "notuserdef" R + val R = remove_junk call_str call_f + ["local", "schematic", "nocompute", "unlisted"] + R + val _ = null R orelse + raise mk_HOL_ERR "boolSyntax" call_str + ("Unhandled reserved attribute(s): " ^ + String.concatWith ", " (map #1 R)) val attrs = if notuserdefp orelse not (is_attribute "userdef") orelse Theory.is_temp_binding s0 then attrs - else "userdef" :: attrs + else ("userdef",[]) :: attrs val th = genth (s0, arg) - fun do_attr a = store_at_attribute {thm = th, name = s0, attrname = a} + fun do_attr (k,vs) = + store_at_attribute {thm = th, name = s0, attrname = k, args = vs} in List.app do_attr attrs; th end diff --git a/src/1/selftest.sml b/src/1/selftest.sml index 47e3df2803..7e1f706c5c 100644 --- a/src/1/selftest.sml +++ b/src/1/selftest.sml @@ -1315,3 +1315,29 @@ in SWAP_EXISTS_CONV, t, expected) end + +fun q s = "\"" ^ String.toString s ^ "\"" +fun kvs_toString (k,vs) = + "(" ^ q k ^ ", [" ^ String.concatWith ", " (map q vs) ^ "])" +fun kvs_alist_toString al = + "[" ^ String.concatWith ", " (map kvs_toString al) ^ "]" +fun attr_result_toString {thmname,attrs,reserved,unknown} = + "{ attrs = " ^ kvs_alist_toString attrs ^ ",\n" ^ + " reserved = " ^ kvs_alist_toString reserved ^ ",\n" ^ + " unknown = " ^ kvs_alist_toString unknown ^ ",\n" ^ + " thmname = " ^ q thmname ^ "}" + +val _ = let + val _ = tprint ("extract_attributes \"" ^ + "foo[local,simp=once twice,baz]\"") + val expected = {attrs = [] : (string * string list) list, + reserved = [("local", [])], + thmname = "foo", + unknown = [("simp", ["once", "twice"]), ("baz", [])]} +in + require_msg + (check_result (equal expected)) + attr_result_toString + ThmAttribute.extract_attributes + "foo[local,simp=once twice,baz]" +end diff --git a/src/1/theory_tests/Holmakefile b/src/1/theory_tests/Holmakefile new file mode 100644 index 0000000000..2d13e485b2 --- /dev/null +++ b/src/1/theory_tests/Holmakefile @@ -0,0 +1,3 @@ +ifdef POLY +CLINE_OPTIONS = --holstate=$(HOLDIR)/bin/hol.state0 +endif diff --git a/src/1/theory_tests/gh1370Script.sml b/src/1/theory_tests/gh1370Script.sml new file mode 100644 index 0000000000..7a1aa2a208 --- /dev/null +++ b/src/1/theory_tests/gh1370Script.sml @@ -0,0 +1,15 @@ +open HolKernel boolLib + +val _ = new_theory"gh1370"; + +val _ = new_definition ("def", “c = T”); + +val _ = + adjoin_to_theory + {sig_ps = SOME (fn _ => PP.add_string "val ctm : Term.term"), + struct_ps = NONE}; + +val _ = adjoin_after_completion + (fn _ => PP.add_string ("val ctm = Parse.Term ‘c’;\n\n")) + +val _ = export_theory(); diff --git a/src/1/theory_tests/gh1370childScript.sml b/src/1/theory_tests/gh1370childScript.sml new file mode 100644 index 0000000000..f7493ac403 --- /dev/null +++ b/src/1/theory_tests/gh1370childScript.sml @@ -0,0 +1,11 @@ +open HolKernel Parse boolLib + +open gh1370Theory + +val _ = new_theory "gh1370child"; + +val _ = type_of gh1370Theory.ctm = bool andalso + (print "ctm has correct type\n";true) orelse + OS.Process.exit OS.Process.failure + +val _ = export_theory(); diff --git a/src/bag/primeFactorScript.sml b/src/bag/primeFactorScript.sml index 2e34ccbd0a..1ecc255df6 100644 --- a/src/bag/primeFactorScript.sml +++ b/src/bag/primeFactorScript.sml @@ -8,7 +8,6 @@ open simpLib BasicProvers metisLib bagTheory dividesTheory arithmeticTheory; val _ = new_theory "primeFactor"; -infix ++; val std_ss = (boolSimps.bool_ss ++ pairSimps.PAIR_ss ++ optionSimps.OPTION_ss ++ numSimps.REDUCE_ss ++ sumSimps.SUM_ss ++ combinSimps.COMBIN_ss ++ diff --git a/src/bool/boolScript.sml b/src/bool/boolScript.sml index 09f5e3a950..6bd1019e50 100644 --- a/src/bool/boolScript.sml +++ b/src/bool/boolScript.sml @@ -293,7 +293,6 @@ val DATATYPE_TAG_DEF = def (#(FILE), #(LINE)) *---------------------------------------------------------------------------*) val op --> = Type.--> -infix ## |->; val ERR = Feedback.mk_HOL_ERR "boolScript" diff --git a/src/boss/prove_base_assumsScript.sml b/src/boss/prove_base_assumsScript.sml index 3d4fded15b..773521713f 100644 --- a/src/boss/prove_base_assumsScript.sml +++ b/src/boss/prove_base_assumsScript.sml @@ -296,16 +296,17 @@ val goals = 60 |- (!P. Data_List_any P Data_List_nil <=> F) /\ !P h t. Data_List_any P (Data_List_cons h t) <=> P h \/ Data_List_any P t, - 61 |- Data_List_concat Data_List_nil = Data_List_nil /\ + 61 |- (T <> F) /\ (F <> T) + 62 |- Data_List_concat Data_List_nil = Data_List_nil /\ !h t. Data_List_concat (Data_List_cons h t) = Data_List_append h (Data_List_concat t), - 62 |- Data_List_reverse Data_List_nil = Data_List_nil /\ + 63 |- Data_List_reverse Data_List_nil = Data_List_nil /\ !h t. Data_List_reverse (Data_List_cons h t) = Data_List_append (Data_List_reverse t) (Data_List_cons h Data_List_nil), - 63 |- Data_List_unzip Data_List_nil = + 64 |- Data_List_unzip Data_List_nil = Data_Pair_comma Data_List_nil Data_List_nil /\ !x l. Data_List_unzip (Data_List_cons x l) = @@ -314,59 +315,59 @@ val goals = (Data_Pair_fst (Data_List_unzip l))) (Data_List_cons (Data_Pair_snd x) (Data_Pair_snd (Data_List_unzip l))), - 64 |- Data_List_length Data_List_nil = Number_Natural_zero /\ + 65 |- Data_List_length Data_List_nil = Number_Natural_zero /\ !h t. Data_List_length (Data_List_cons h t) = Number_Natural_suc (Data_List_length t), - 65 |- Number_Natural_factorial Number_Natural_zero = + 66 |- Number_Natural_factorial Number_Natural_zero = Number_Natural_bit1 Number_Natural_zero /\ !n. Number_Natural_factorial (Number_Natural_suc n) = Number_Natural_times (Number_Natural_suc n) (Number_Natural_factorial n), - 66 |- (Data_List_null Data_List_nil <=> T) /\ + 67 |- (Data_List_null Data_List_nil <=> T) /\ !h t. Data_List_null (Data_List_cons h t) <=> F, - 67 |- (Number_Natural_even Number_Natural_zero <=> T) /\ + 68 |- (Number_Natural_even Number_Natural_zero <=> T) /\ !n. Number_Natural_even (Number_Natural_suc n) <=> ~Number_Natural_even n, - 68 |- (Number_Natural_odd Number_Natural_zero <=> F) /\ + 69 |- (Number_Natural_odd Number_Natural_zero <=> F) /\ !n. Number_Natural_odd (Number_Natural_suc n) <=> ~Number_Natural_odd n, - 69 |- Data_Unit_unit = @x. T, - 70 |- Number_Natural_zero = Number_Natural_zero, - 71 |- (?!x. F) <=> F, - 72 |- Data_Pair_comma x y = Data_Pair_comma a b <=> x = a /\ y = b, - 73 |- Data_Sum_left x = Data_Sum_left y <=> x = y, - 74 |- Data_Sum_right x = Data_Sum_right y <=> x = y, - 75 |- Function_id = Function_Combinator_s Function_const Function_const, - 76 |- Relation_empty = (\x y. F), - 77 |- Relation_universe = (\x y. T), - 78 |- Number_Natural_bit1 = + 70 |- Data_Unit_unit = @x. T, + 71 |- Number_Natural_zero = Number_Natural_zero, + 72 |- (?!x. F) <=> F, + 73 |- Data_Pair_comma x y = Data_Pair_comma a b <=> x = a /\ y = b, + 74 |- Data_Sum_left x = Data_Sum_left y <=> x = y, + 75 |- Data_Sum_right x = Data_Sum_right y <=> x = y, + 76 |- Function_id = Function_Combinator_s Function_const Function_const, + 77 |- Relation_empty = (\x y. F), + 78 |- Relation_universe = (\x y. T), + 79 |- Number_Natural_bit1 = (\n. Number_Natural_plus n (Number_Natural_plus n (Number_Natural_suc Number_Natural_zero))), - 79 |- Number_Natural_max = (\m n. if Number_Natural_less m n then n else m), - 80 |- Number_Natural_min = (\m n. if Number_Natural_less m n then m else n), - 81 |- Number_Natural_greater = (\m n. Number_Natural_less n m), - 82 |- Number_Natural_greatereq = (\m n. Number_Natural_greater m n \/ m = n), - 83 |- Number_Natural_less = + 80 |- Number_Natural_max = (\m n. if Number_Natural_less m n then n else m), + 81 |- Number_Natural_min = (\m n. if Number_Natural_less m n then m else n), + 82 |- Number_Natural_greater = (\m n. Number_Natural_less n m), + 83 |- Number_Natural_greatereq = (\m n. Number_Natural_greater m n \/ m = n), + 84 |- Number_Natural_less = (\m n. ?P. (!n. P (Number_Natural_suc n) ==> P n) /\ P m /\ ~P n), - 84 |- Number_Natural_lesseq = (\m n. Number_Natural_less m n \/ m = n), - 85 |- Relation_irreflexive = (\R. !x. ~R x x), - 86 |- Relation_reflexive = (\R. !x. R x x), - 87 |- Relation_transitive = (\R. !x y z. R x y /\ R y z ==> R x z), - 88 |- Relation_wellFounded = + 85 |- Number_Natural_lesseq = (\m n. Number_Natural_less m n \/ m = n), + 86 |- Relation_irreflexive = (\R. !x. ~R x x), + 87 |- Relation_reflexive = (\R. !x. R x x), + 88 |- Relation_transitive = (\R. !x y z. R x y /\ R y z ==> R x z), + 89 |- Relation_wellFounded = (\R. !B. (?w. B w) ==> ?min. B min /\ !b. R b min ==> ~B b), - 89 |- Relation_transitiveClosure = + 90 |- Relation_transitiveClosure = (\R a b. !P. (!x y. R x y ==> P x y) /\ (!x y z. P x y /\ P y z ==> P x z) ==> P a b), - 90 |- Relation_subrelation = (\R1 R2. !x y. R1 x y ==> R2 x y), - 91 |- Relation_intersect = (\R1 R2 x y. R1 x y /\ R2 x y), - 92 |- Relation_union = (\R1 R2 x y. R1 x y \/ R2 x y), - 93 |- Function_o = (\f g x. f (g x)), - 94 |- (!x. P x ==> Q x) ==> (?x. P x) ==> ?x. Q x, - 95 |- (x ==> y) /\ (z ==> w) ==> x /\ z ==> y /\ w, - 96 |- (x ==> y) /\ (z ==> w) ==> x \/ z ==> y \/ w]: thm list + 91 |- Relation_subrelation = (\R1 R2. !x y. R1 x y ==> R2 x y), + 92 |- Relation_intersect = (\R1 R2 x y. R1 x y /\ R2 x y), + 93 |- Relation_union = (\R1 R2 x y. R1 x y \/ R2 x y), + 94 |- Function_o = (\f g x. f (g x)), + 95 |- (!x. P x ==> Q x) ==> (?x. P x) ==> ?x. Q x, + 96 |- (x ==> y) /\ (z ==> w) ==> x /\ z ==> y /\ w, + 97 |- (x ==> y) /\ (z ==> w) ==> x \/ z ==> y \/ w]: thm list *) val bool_cases = hd(amatch``(x = T) \/ _``); @@ -1085,13 +1086,18 @@ val th60 = store_thm val concat_nil = hd(amatch``Data_List_concat Data_List_nil``); val concat_cons = hd(amatch``Data_List_concat (Data_List_cons _ _)``); +(* |- (T <> F) /\ (F <> T) *) +val th61 = store_thm + ("th61", el 61 goals |> concl, + PURE_REWRITE_TAC[iff_F,not_not,iff_T,not_F,and_T]); + (* |- Data_List_concat Data_List_nil = Data_List_nil /\ !h t. Data_List_concat (Data_List_cons h t) = Data_List_append h (Data_List_concat t) *) -val th61 = store_thm - ("th61", el 61 goals |> concl, +val th62 = store_thm + ("th62", el 62 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC concat_nil \\ MATCH_ACCEPT_TAC concat_cons); @@ -1104,8 +1110,8 @@ val reverse_cons = hd(amatch``Data_List_reverse (Data_List_cons _ _)``); Data_List_append (Data_List_reverse t) (Data_List_cons h Data_List_nil) *) -val th62 = store_thm - ("th62", el 62 goals |> concl, +val th63 = store_thm + ("th63", el 63 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC reverse_nil \\ MATCH_ACCEPT_TAC reverse_cons); @@ -1123,8 +1129,8 @@ val unzip_cons = hd(amatch``Data_List_unzip (Data_List_cons _ _)``); (Data_List_cons (Data_Pair_snd x) (Data_Pair_snd (Data_List_unzip l))) *) -val th63 = store_thm - ("th63", el 63 goals |> concl, +val th64 = store_thm + ("th64", el 64 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC unzip_nil \\ PURE_REWRITE_TAC[unzip_cons] \\ rpt gen_tac @@ -1139,8 +1145,8 @@ val length_cons = hd(amatch``Data_List_length (Data_List_cons _ _)``); Data_List_length (Data_List_cons h t) = Number_Natural_suc (Data_List_length t) *) -val th64 = store_thm - ("th64", el 64 goals |> concl, +val th65 = store_thm + ("th65", el 65 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC length_nil \\ MATCH_ACCEPT_TAC length_cons); @@ -1154,8 +1160,8 @@ val fact_suc = hd(amatch``Number_Natural_factorial (Number_Natural_suc _)``); Number_Natural_times (Number_Natural_suc n) (Number_Natural_factorial n) *) -val th65 = store_thm - ("th65", el 65 goals |> concl, +val th66 = store_thm + ("th66", el 66 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC fact_zero \\ MATCH_ACCEPT_TAC fact_suc); @@ -1165,8 +1171,8 @@ val null_cons = hd(amatch``Data_List_null (Data_List_cons _ _)``); (* |- (Data_List_null Data_List_nil <=> T) /\ !h t. Data_List_null (Data_List_cons h t) <=> F *) -val th66 = store_thm - ("th66", el 66 goals |> concl, +val th67 = store_thm + ("th67", el 67 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC (EQT_INTRO null_nil) \\ MATCH_ACCEPT_TAC (EQF_INTRO (SPEC_ALL null_cons))); @@ -1176,8 +1182,8 @@ val even_cons = hd(amatch``Number_Natural_even (Number_Natural_suc _)``); (* |- (Number_Natural_even Number_Natural_zero <=> T) /\ !n. Number_Natural_even (Number_Natural_suc n) <=> ~Number_Natural_even n *) -val th67 = store_thm - ("th67", el 67 goals |> concl, +val th68 = store_thm + ("th68", el 68 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC (EQT_INTRO even_nil) \\ MATCH_ACCEPT_TAC even_cons); @@ -1187,28 +1193,28 @@ val odd_cons = hd(amatch``Number_Natural_odd (Number_Natural_suc _)``); (* |- (Number_Natural_odd Number_Natural_zero <=> F) /\ !n. Number_Natural_odd (Number_Natural_suc n) <=> ~Number_Natural_odd n *) -val th68 = store_thm - ("th68", el 68 goals |> concl, +val th69 = store_thm + ("th69", el 69 goals |> concl, conj_tac >- MATCH_ACCEPT_TAC (EQF_INTRO odd_nil) \\ MATCH_ACCEPT_TAC odd_cons); val one_thm = hd(amatch``_ = Data_Unit_unit``); (* |- Data_Unit_unit = @x. T *) -val th69 = store_thm - ("th69", el 69 goals |> concl, +val th70 = store_thm + ("th70", el 70 goals |> concl, PURE_ONCE_REWRITE_TAC[one_thm] \\ REFL_TAC); (* |- Number_Natural_zero = Number_Natural_zero *) -val th70 = store_thm - ("th70", el 70 goals |> concl, +val th71 = store_thm + ("th71", el 71 goals |> concl, REFL_TAC); val exists_simp = hd(amatch “(?x. t) <=> t”); (* |- (?!x. F) <=> F *) -val th71 = store_thm - ("th71", el 71 goals |> concl, +val th72 = store_thm + ("th72", el 72 goals |> concl, PURE_REWRITE_TAC [BETA_RULE (SPEC “\x:'a. F” ex_unique_thm)] \\ PURE_REWRITE_TAC [SPEC “F” exists_simp, F_and] \\ REFL_TAC); @@ -1216,32 +1222,32 @@ val th71 = store_thm val comma_11 = hd(amatch``Data_Pair_comma _ _ = Data_Pair_comma _ _``); (* |- Data_Pair_comma x y = Data_Pair_comma a b <=> x = a /\ y = b *) -val th72 = store_thm - ("th72", el 72 goals |> concl, +val th73 = store_thm + ("th73", el 73 goals |> concl, MATCH_ACCEPT_TAC comma_11); (* |- Data_Sum_left x = Data_Sum_left y <=> x = y *) -val th73 = store_thm - ("th73", el 73 goals |> concl, +val th74 = store_thm + ("th74", el 74 goals |> concl, MATCH_ACCEPT_TAC left_11); (* |- Data_Sum_right x = Data_Sum_right y <=> x = y *) -val th74 = store_thm - ("th74", el 74 goals |> concl, +val th75 = store_thm + ("th75", el 75 goals |> concl, MATCH_ACCEPT_TAC right_11); val skk = hd(amatch``Function_Combinator_s _ _ = Function_id``); (* |- Function_id = Function_Combinator_s Function_const Function_const *) -val th75 = store_thm - ("th75", el 75 goals |> concl, +val th76 = store_thm + ("th76", el 76 goals |> concl, PURE_REWRITE_TAC[skk] \\ REFL_TAC); val empty_thm = hd(amatch``Relation_empty _ _``); (* |- Relation_empty = (\x y. F) *) -val th76 = store_thm - ("th76", el 76 goals |> concl, +val th77 = store_thm + ("th77", el 77 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,EQF_INTRO (SPEC_ALL empty_thm)] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1249,8 +1255,8 @@ val th76 = store_thm val universe_thm = hd(amatch``Relation_universe _ _``); (* |- Relation_universe = (\x y. T) *) -val th77 = store_thm - ("th77", el 77 goals |> concl, +val th78 = store_thm + ("th78", el 78 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,EQT_INTRO(SPEC_ALL universe_thm)] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1271,8 +1277,8 @@ val num_less_ind = hd(amatch``(!x. _ ==> _) ==> !n. P (n:Number_Natural_natural) Number_Natural_plus n (Number_Natural_plus n (Number_Natural_suc Number_Natural_zero))) *) -val th78 = store_thm - ("th78", el 78 goals |> concl, +val th79 = store_thm + ("th79", el 79 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,bit1_thm,plus_suc] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ gen_tac @@ -1291,8 +1297,8 @@ val if_id = hd(amatch``if _ then x else x``); val less_or_eq = hd(amatch``Number_Natural_lesseq _ _ <=> (Number_Natural_less _ _) \/ _``); (* |- Number_Natural_max = (\m n. if Number_Natural_less m n then n else m) *) -val th79 = store_thm - ("th79", el 79 goals |> concl, +val th80 = store_thm + ("th80", el 80 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,max_thm,less_or_eq] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac @@ -1307,8 +1313,8 @@ val th79 = store_thm val min_thm = hd(amatch``Number_Natural_min _ _ = COND _ _ _``); (* |- Number_Natural_min = (\m n. if Number_Natural_less m n then m else n) *) -val th80 = store_thm - ("th80", el 80 goals |> concl, +val th81 = store_thm + ("th81", el 81 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,min_thm,less_or_eq] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac @@ -1323,8 +1329,8 @@ val th80 = store_thm val greater_thm = hd(amatch``Number_Natural_greater _ _ = _``); (* |- Number_Natural_greater = (\m n. Number_Natural_less n m) *) -val th81 = store_thm - ("th81", el 81 goals |> concl, +val th82 = store_thm + ("th82", el 82 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,greater_thm] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1332,8 +1338,8 @@ val th81 = store_thm val greatereq_thm = hd(amatch``Number_Natural_greatereq _ _``); (* |- Number_Natural_greatereq = (\m n. Number_Natural_greater m n \/ m = n) *) -val th82 = store_thm - ("th82", el 82 goals |> concl, +val th83 = store_thm + ("th83", el 83 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,greatereq_thm,less_or_eq,greater_thm] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac @@ -1361,8 +1367,8 @@ val F_IMP2 = store_thm (* |- Number_Natural_less = (\m n. ?P. (!n. P (Number_Natural_suc n) ==> P n) /\ P m /\ ~P n) *) -val th83 = store_thm - ("th83", el 83 goals |> concl, +val th84 = store_thm + ("th84", el 84 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm] \\ qx_genl_tac[`a`,`b`] \\ CONV_TAC(DEPTH_CONV BETA_CONV) @@ -1421,8 +1427,8 @@ val th83 = store_thm \\ PURE_REWRITE_TAC[F_imp]); (* |- Number_Natural_lesseq = (\m n. Number_Natural_less m n \/ m = n) *) -val th84 = store_thm - ("th84", el 84 goals |> concl, +val th85 = store_thm + ("th85", el 85 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,less_or_eq] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1430,8 +1436,8 @@ val th84 = store_thm val irreflexive_thm = hd(amatch``Relation_irreflexive _ = _``); (* |- Relation_irreflexive = (\R. !x. ~R x x) *) -val th85 = store_thm - ("th85", el 85 goals |> concl, +val th86 = store_thm + ("th86", el 86 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm, irreflexive_thm] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac @@ -1440,8 +1446,8 @@ val th85 = store_thm val reflexive_thm = hd(amatch``Relation_reflexive _ = _``); (* |- Relation_reflexive = (\R. !x. R x x) *) -val th86 = store_thm - ("th86", el 86 goals |> concl, +val th87 = store_thm + ("th87", el 87 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm, reflexive_thm] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1449,8 +1455,8 @@ val th86 = store_thm val transitive_thm = hd(amatch``Relation_transitive s <=> _``); (* |- Relation_transitive = (\R. !x y z. R x y /\ R y z ==> R x z) *) -val th87 = store_thm - ("th87", el 87 goals |> concl, +val th88 = store_thm + ("th88", el 88 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,transitive_thm] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ gen_tac \\ REFL_TAC); @@ -1460,8 +1466,8 @@ val wellFounded_thm = hd(amatch``Relation_wellFounded r <=> !p. (?x. _) ==> _``) (* |- Relation_wellFounded = (\R. !B. (?w. B w) ==> ?min. B min /\ !b. R b min ==> ~B b) *) -val th88 = store_thm - ("th88", el 88 goals |> concl, +val th89 = store_thm + ("th89", el 89 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm] \\ PURE_REWRITE_TAC[wellFounded_thm] \\ CONV_TAC(DEPTH_CONV BETA_CONV) @@ -1477,8 +1483,8 @@ val subrelation_thm = hd(amatch``Relation_subrelation x s <=> !x y. _``); !P. (!x y. R x y ==> P x y) /\ (!x y z. P x y /\ P y z ==> P x z) ==> P a b) *) -val th89 = store_thm - ("th89", el 89 goals |> concl, +val th90 = store_thm + ("th90", el 90 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm] \\ PURE_REWRITE_TAC[tc_def,bigIntersect_thm,mem_fromPred] \\ CONV_TAC(DEPTH_CONV BETA_CONV) @@ -1491,8 +1497,8 @@ val th89 = store_thm \\ REFL_TAC); (* |- Relation_subrelation = (\R1 R2. !x y. R1 x y ==> R2 x y) *) -val th90 = store_thm - ("th90", el 90 goals |> concl, +val th91 = store_thm + ("th91", el 91 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,subrelation_thm] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1505,15 +1511,15 @@ val mem_union = hd(amatch``Set_member _ (Set_union _ _) <=> _``); val mem_toSet = hd(amatch``Set_member _ (Relation_toSet _)``); (* |- Relation_intersect = (\R1 R2 x y. R1 x y /\ R2 x y) *) -val th91 = store_thm - ("th91", el 91 goals |> concl, +val th92 = store_thm + ("th92", el 92 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,intersect_thm,fromSet_thm,mem_inter,mem_toSet] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); (* |- Relation_union = (\R1 R2 x y. R1 x y \/ R2 x y) *) -val th92 = store_thm - ("th92", el 92 goals |> concl, +val th93 = store_thm + ("th93", el 93 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,union_thm,fromSet_thm,mem_union,mem_toSet] \\ CONV_TAC(DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); @@ -1521,15 +1527,15 @@ val th92 = store_thm val o_thm = hd(amatch``(Function_o _ _) _ = _``); (* |- Function_o = (\f g x. f (g x)) *) -val th93 = store_thm - ("th93", el 93 goals |> concl, +val th94 = store_thm + ("th94", el 94 goals |> concl, PURE_REWRITE_TAC[GSYM fun_eq_thm,o_thm] \\ CONV_TAC (DEPTH_CONV BETA_CONV) \\ rpt gen_tac \\ REFL_TAC); (* |- (!x. P x ==> Q x) ==> (?x. P x) ==> ?x'. Q x' *) -val th94 = store_thm - ("th94", el 94 goals |> concl, +val th95 = store_thm + ("th95", el 95 goals |> concl, rpt strip_tac \\ first_x_assum(fn th => first_x_assum (assume_tac o MATCH_MP th)) \\ qexists_tac`x` @@ -1544,16 +1550,16 @@ val th94' = store_thm *) (* |- (x ==> y) /\ (z ==> w) ==> x /\ z ==> y /\ w *) -val th95 = store_thm - ("th95", el 95 goals |> concl, +val th96 = store_thm + ("th96", el 96 goals |> concl, rpt strip_tac \\ first_x_assum(fn th => first_x_assum (assume_tac o MATCH_MP th)) \\ first_x_assum(fn th => first_x_assum (assume_tac o MATCH_MP th)) \\ first_assum ACCEPT_TAC); (* |- (x ==> y) /\ (z ==> w) ==> x \/ z ==> y \/ w *) -val th96 = store_thm - ("th96", el 96 goals |> concl, +val th97 = store_thm + ("th97", el 97 goals |> concl, rpt strip_tac \\ first_x_assum(fn th => first_x_assum (assume_tac o MATCH_MP th)) \\ TRY (disj1_tac >> first_assum ACCEPT_TAC) @@ -1590,7 +1596,7 @@ val th89 = store_thm *) (* Now raise an error if the above th96 is not the last one *) -val _ = if List.length goals <> 96 then +val _ = if List.length goals <> 97 then (raise ERR "-" "assumptions changed") else (); @@ -1629,8 +1635,7 @@ val LET_RAND = store_thm("LET_RAND", concl boolTheory.LET_RAND, \\ REFL_TAC) (* |- (T <=/=> F) /\ (F <=/=> T) *) -val BOOL_EQ_DISTINCT = store_thm("BOOL_EQ_DISTINCT", concl boolTheory.BOOL_EQ_DISTINCT, - PURE_REWRITE_TAC[iff_F,not_not,iff_T,not_F,and_T]); +val BOOL_EQ_DISTINCT = th61; (* |- (!t1 t2. (if T then t1 else t2) = t1) /\ !t1 t2. (if F then t1 else t2) = t2 diff --git a/src/coalgebras/ltreeScript.sml b/src/coalgebras/ltreeScript.sml index 5179b3264e..9f10b18439 100644 --- a/src/coalgebras/ltreeScript.sml +++ b/src/coalgebras/ltreeScript.sml @@ -879,7 +879,7 @@ Proof QED Theorem ltree_paths_alt : - !t. ltree_paths A = {p | ltree_el A p <> NONE} + !t. ltree_paths t = {p | ltree_el t p <> NONE} Proof rw [ltree_paths_def, Once EXTENSION, ltree_lookup_iff_ltree_el] QED diff --git a/src/coalgebras/pathScript.sml b/src/coalgebras/pathScript.sml index d32e64634b..47e86002e2 100644 --- a/src/coalgebras/pathScript.sml +++ b/src/coalgebras/pathScript.sml @@ -1139,9 +1139,6 @@ val okpath_monotone = store_thm( val okpath_def = Define`okpath R = gfp (okpath_f R)`; -infix |> -fun x |> f = f x - val okpath_co_ind = save_thm( "okpath_co_ind", okpath_monotone |> SPEC_ALL diff --git a/src/combin/combinScript.sml b/src/combin/combinScript.sml index a829a8752b..e2a8880f5e 100644 --- a/src/combin/combinScript.sml +++ b/src/combin/combinScript.sml @@ -10,7 +10,7 @@ (* AUGMENTED : (kxs) added C and W combinators *) (* ===================================================================== *) -open HolKernel Parse boolLib; +open HolKernel Parse boolLib computeLib; val _ = new_theory "combin"; @@ -20,14 +20,14 @@ val _ = new_theory "combin"; fun def (s,l) p = Q.new_definition_at (DB_dtype.mkloc(s,l,false)) p -val K_DEF = def(#(FILE),#(LINE))("K_DEF", `K = \x y. x`); -val S_DEF = def(#(FILE),#(LINE))("S_DEF", `S = \f g x. f x (g x)`); -val I_DEF = def(#(FILE),#(LINE))("I_DEF", `I = S K (K:'a->'a->'a)`); -val C_DEF = def(#(FILE),#(LINE))("C_DEF", `C = \f x y. f y x`); -val W_DEF = def(#(FILE),#(LINE))("W_DEF", `W = \f x. f x x`); -val o_DEF = def(#(FILE),#(LINE))("o_DEF", `$o f g = \x. f(g x)`); +val K_DEF = def(#(FILE),#(LINE))("K_DEF", ‘K = \x y. x’); +val S_DEF = def(#(FILE),#(LINE))("S_DEF[compute]", ‘S = \f g x. f x (g x)’); +val I_DEF = def(#(FILE),#(LINE))("I_DEF", ‘I = S K (K:'a->'a->'a)’); +val C_DEF = def(#(FILE),#(LINE))("C_DEF[compute]", ‘C = \f x y. f y x’); +val W_DEF = def(#(FILE),#(LINE))("W_DEF[compute]", ‘W = \f x. f x x’); +val o_DEF = def(#(FILE),#(LINE))("o_DEF", ‘$o f g = \x. f(g x)’); val _ = set_fixity "o" (Infixr 800) -val APP_DEF = def(#(FILE),#(LINE)) ("APP_DEF", `$:> x f = f x`); +val APP_DEF = def(#(FILE),#(LINE))("APP_DEF[compute]",‘$:> x f = f x’); val UPDATE_def = def(#(FILE),#(LINE))("UPDATE_def", `UPDATE a b = \f c. if a = c then b else f c`); @@ -92,12 +92,14 @@ end * superset of those on the rhs. * *---------------------------------------------------------------------------*) -val o_THM = store_thm("o_THM", - “!f g x. (f o g) x = f(g x)”, +Theorem o_THM[compute]: + !f g x. (f o g) x = f(g x) +Proof REPEAT GEN_TAC THEN PURE_REWRITE_TAC [ o_DEF ] THEN CONV_TAC (DEPTH_CONV BETA_CONV) - THEN REFL_TAC); + THEN REFL_TAC +QED val o_ASSOC = store_thm("o_ASSOC", “!f g h. f o (g o h) = (f o g) o h”, @@ -118,12 +120,14 @@ val o_ABS_R = store_thm( ``f o (\x. g x) = (\x. f (g x))``, REWRITE_TAC [FUN_EQ_THM, o_THM] THEN BETA_TAC THEN REWRITE_TAC []); -val K_THM = store_thm("K_THM", - “!x y. K x y = x”, +Theorem K_THM[compute]: + !x y. K x y = x +Proof REPEAT GEN_TAC THEN PURE_REWRITE_TAC [ K_DEF ] THEN CONV_TAC (DEPTH_CONV BETA_CONV) - THEN REFL_TAC); + THEN REFL_TAC +QED Theorem K_PARTIAL : (* from seqTheory *) !x. K x = \z. x @@ -168,12 +172,14 @@ val W_THM = store_thm("W_THM", THEN CONV_TAC (DEPTH_CONV BETA_CONV) THEN REFL_TAC); -val I_THM = store_thm("I_THM", - “!x. I x = x”, +Theorem I_THM[compute]: + !x. I x = x +Proof REPEAT GEN_TAC THEN PURE_REWRITE_TAC [ I_DEF, S_THM, K_THM ] THEN CONV_TAC (DEPTH_CONV BETA_CONV) - THEN REFL_TAC); + THEN REFL_TAC +QED Theorem I_EQ_IDABS: I = \x. x @@ -185,9 +191,11 @@ val I_o_ID = store_thm("I_o_ID", “!f. (I o f = f) /\ (f o I = f)”, REWRITE_TAC [I_THM, o_THM, FUN_EQ_THM]); -val K_o_THM = store_thm("K_o_THM", - “(!f v. K v o f = K v) /\ (!f v. f o K v = K (f v))”, - REWRITE_TAC [o_THM, K_THM, FUN_EQ_THM]); +Theorem K_o_THM[compute]: + (!f v. K v o f = K v) /\ (!f v. f o K v = K (f v)) +Proof + REWRITE_TAC [o_THM, K_THM, FUN_EQ_THM] +QED val UPDATE_APPLY = Q.store_thm("UPDATE_APPLY", `(!a x f. (a =+ x) f a = x) /\ @@ -200,10 +208,12 @@ val UPDATE_APPLY = Q.store_thm("UPDATE_APPLY", Theorem UPDATE_APPLY1 = cj 1 UPDATE_APPLY -val APPLY_UPDATE_THM = Q.store_thm("APPLY_UPDATE_THM", - `!f a b c. (a =+ b) f c = (if a = c then b else f c)`, +Theorem APPLY_UPDATE_THM[compute]: + !f a b c. (a =+ b) f c = (if a = c then b else f c) +Proof PURE_REWRITE_TAC [UPDATE_def] - THEN BETA_TAC THEN REWRITE_TAC []); + THEN BETA_TAC THEN REWRITE_TAC [] +QED val UPDATE_COMMUTES = Q.store_thm("UPDATE_COMMUTES", `!f a b c d. ~(a = b) ==> ((a =+ c) ((b =+ d) f) = (b =+ d) ((a =+ c) f))`, @@ -553,37 +563,4 @@ Proof >> ASM_REWRITE_TAC [] QED -(*---------------------------------------------------------------------------*) - -val _ = adjoin_to_theory -{sig_ps = NONE, - struct_ps = SOME (fn _ => - let val S = PP.add_string fun SPC n = PP.add_break(1,n) - fun B n = PP.block PP.CONSISTENT n - fun I n = PP.block PP.INCONSISTENT n - val L = PP.pr_list - in - B 0 [ - S "val _ =", SPC 2, - B 4 [ - S "let open computeLib", SPC 0, - B 2 [ - S "val K_tm =", SPC 0, - S "Term.prim_mk_const{Name=\"K\",Thy=\"combin\"}" - ], SPC ~4, - S "in", SPC 0, - B 2 [ - S "add_funs", SPC 0, - I 1 (S "[" :: - L S [S ",", PP.add_break(0,0)] [ - "K_THM", "S_DEF", "I_THM", "C_DEF", "W_DEF", "o_THM", - "K_o_THM", "APP_DEF", "APPLY_UPDATE_THM];" - ]) - ], SPC 0, - B 2 (L S [SPC 1] ["set_skip", "the_compset", "K_tm", "(SOME 1)"]), - SPC ~4, S "end;" - ] - ] - end)} - val _ = export_theory(); diff --git a/src/compute/src/computeLib.sig b/src/compute/src/computeLib.sig index ff52cede03..f99a491b9f 100644 --- a/src/compute/src/computeLib.sig +++ b/src/compute/src/computeLib.sig @@ -15,6 +15,8 @@ sig val add_conv : term * int * conv -> compset -> unit val add_thmset : string -> compset -> unit val set_skip : compset -> term -> int option -> unit + val set_EVAL_skip : term -> int option -> unit + val temp_set_EVAL_skip : term -> int option -> unit val scrub_const : compset -> term -> unit val scrub_thms : thm list -> compset -> unit diff --git a/src/compute/src/computeLib.sml b/src/compute/src/computeLib.sml index 0977bddaea..76d83fcdc0 100644 --- a/src/compute/src/computeLib.sml +++ b/src/compute/src/computeLib.sml @@ -392,38 +392,53 @@ val {export,...} = } val add_persistent_funs = app export -(*---------------------------------------------------------------------------*) -(* Once we delete a constant from a compset, we will probably want to make *) -(* sure that the constant doesn't get re-added when the theory is exported *) -(*---------------------------------------------------------------------------*) -fun pr_list_to_ppstream pps f b1 b2 [] = () - | pr_list_to_ppstream pps f b1 b2 [e] = f pps e - | pr_list_to_ppstream pps f b1 b2 (e::es) = - (f pps e; b1 pps; b2 pps; pr_list_to_ppstream pps f b1 b2 es) - -fun del_persistent_consts [] = () - | del_persistent_consts clist = - let open Portable - val plist = map (fn c => let val {Name,Thy,Ty} = dest_thy_const c - in (Name,Thy) end) clist - val plist' = map (Lib.mlquote##Lib.mlquote) plist - fun prec (s1,s2) = "Term.prim_mk_const{Name = "^s1^", Thy = "^s2^"}" - val plist'' = map prec plist' - in - del_consts clist - ; Theory.adjoin_to_theory - {sig_ps = NONE, - struct_ps = SOME(fn _ => - PP.block CONSISTENT 0 [ - PP.add_string "val _ = computeLib.del_consts [", - PP.block INCONSISTENT 0 ( - PP.pr_list PP.add_string - [PP.add_string ",", PP.add_break (0,0)] plist'' - ), - PP.add_string "];" - ])} - end +(* ---------------------------------------------------------------------- + Other exported changes: set_skip and del_consts + ---------------------------------------------------------------------- *) + +datatype clib_delta = CLD_set_skip of term * int option + | CLD_delconst of term +fun uptodate_delta (CLD_set_skip(t,_)) = uptodate_term t + | uptodate_delta (CLD_delconst t) = uptodate_term t + +fun encdelta cld = + let open ThyDataSexp + in + case cld of + CLD_set_skip d => pair_encode (Term, option_encode Int) d + | CLD_delconst t => Term t + end + +fun decdelta d = + let open ThyDataSexp + in + case d of + Term t => SOME (CLD_delconst t) + | _ => (pair_decode(term_decode, option_decode int_decode) >> + CLD_set_skip) d + end + +fun apply_delta cld () = + case cld of + CLD_set_skip (t, iopt) => set_skip the_compset t iopt + | CLD_delconst t => del_consts [t] + +val {record_delta,get_deltas=thy_updates,...} = + AncestryData.fullmake { + adinfo = {tag = "computeLib.delta", + initial_values = [("min", ())], + apply_delta = apply_delta}, + uptodate_delta = uptodate_delta, + sexps = {dec = decdelta, enc = encdelta}, + globinfo = {apply_to_global = apply_delta, + thy_finaliser = NONE, + initial_value = ()} + } + +fun set_EVAL_skip t i = record_delta (CLD_set_skip(t,i)) +fun temp_set_EVAL_skip t i = set_skip the_compset t i +fun del_persistent_consts cs = app (fn c => record_delta(CLD_delconst c)) cs (* ---------------------------------------------------------------------- compset pretty-printer diff --git a/src/float/ieeeScript.sml b/src/float/ieeeScript.sml index 6703ba1e75..7f397fc3ed 100644 --- a/src/float/ieeeScript.sml +++ b/src/float/ieeeScript.sml @@ -7,8 +7,6 @@ * First, make standard environment available. * *---------------------------------------------------------------------------*) open HolKernel Parse boolLib; -infixr 3 -->; -infix ## |-> THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL; (*---------------------------------------------------------------------------* * Next, bring in extra tools used. * diff --git a/src/integer/DeepSyntaxScript.sml b/src/integer/DeepSyntaxScript.sml index d2a7111327..4b66c49f00 100644 --- a/src/integer/DeepSyntaxScript.sml +++ b/src/integer/DeepSyntaxScript.sml @@ -1,8 +1,5 @@ open HolKernel Parse boolLib -infix THEN THENL |-> -infix 8 by - open Datatype integerTheory bossLib int_arithTheory simpLib pred_setTheory val _ = new_theory "DeepSyntax"; diff --git a/src/integer/int_arithScript.sml b/src/integer/int_arithScript.sml index a9bf28a137..35dcad606e 100644 --- a/src/integer/int_arithScript.sml +++ b/src/integer/int_arithScript.sml @@ -1026,7 +1026,6 @@ val positive_mod_part = prove( REPEAT STRIP_TAC THEN MATCH_MP_TAC INT_MOD_UNIQUE THEN ASM_SIMP_TAC bool_ss [INT_LT_GT] THEN PROVE_TAC []); -infix 8 on val int_ss = srw_ss() ++ numSimps.ARITH_ss val tac1 = Q.EXISTS_TAC `&n - r` THEN @@ -1035,9 +1034,8 @@ val tac1 = `0 < r` by FULL_SIMP_TAC bool_ss [INT_LE_LT] THEN FULL_SIMP_TAC int_ss [GSYM INT_ADD_ASSOC, INT_SUB_ADD2, less_to_leq_samer, INT_ADD_COMM] THEN - SUBST_ALL_TAC on - (`&n + q * &n = (q + 1) * &n`, - SIMP_TAC bool_ss [INT_ADD_COMM, INT_RDISTRIB, INT_MUL_LID]) THEN + Q.SUBGOAL_THEN `&n + q * &n = (q + 1) * &n` SUBST_ALL_TAC + >- SIMP_TAC bool_ss [INT_ADD_COMM, INT_RDISTRIB, INT_MUL_LID] THEN ASM_SIMP_TAC int_ss [INT_MOD_COMMON_FACTOR] val tac2 = STRIP_TAC THEN REPEAT VAR_EQ_TAC THEN @@ -1047,6 +1045,7 @@ val tac2 = `i < &n` by ASM_SIMP_TAC bool_ss [less_to_leq_samel, GSYM int_sub] THEN ASM_SIMP_TAC bool_ss [positive_mod_part] THEN STRIP_TAC THEN VAR_EQ_TAC THEN FULL_SIMP_TAC int_ss [] + val NOT_INT_DIVIDES = store_thm( "NOT_INT_DIVIDES", ``!c d. ~(c = 0) ==> diff --git a/src/integer/int_bitwiseScript.sml b/src/integer/int_bitwiseScript.sml index c338ab5cff..3aabb67540 100644 --- a/src/integer/int_bitwiseScript.sml +++ b/src/integer/int_bitwiseScript.sml @@ -1,9 +1,6 @@ open HolKernel Parse boolLib bossLib; open integerTheory arithmeticTheory bitTheory intLib listTheory; -infix \\ -val op \\ = op THEN; - val _ = new_theory "int_bitwise"; val _ = ParseExtras.temp_loose_equality() diff --git a/src/list/src/listScript.sml b/src/list/src/listScript.sml index 84a465bb0d..e728fddbe0 100644 --- a/src/list/src/listScript.sml +++ b/src/list/src/listScript.sml @@ -4734,6 +4734,16 @@ val EL_LENGTH_dropWhile_REVERSE = Q.store_thm("EL_LENGTH_dropWhile_REVERSE", >> fs [NOT_EVERY, dropWhile_APPEND_EXISTS, arithmeticTheory.ADD1]) end +Theorem dropWhile_id: + (dropWhile P ls = ls) <=> NULL ls \/ ~P(HD ls) +Proof + Cases_on`ls` \\ rw[dropWhile_def, NULL] + \\ disch_then(mp_tac o Q.AP_TERM`LENGTH`) + \\ Q.MATCH_GOALSUB_RENAME_TAC`dropWhile P l` + \\ Q.SPECL_THEN[`P`,`l`]mp_tac LENGTH_dropWhile_LESS_EQ + \\ simp[] +QED + val IMP_EVERY_LUPDATE = store_thm("IMP_EVERY_LUPDATE", “!xs h i. P h /\ EVERY P xs ==> EVERY P (LUPDATE h i xs)”, Induct THEN fs [LUPDATE_def] THEN REPEAT STRIP_TAC diff --git a/src/list/src/numposrepScript.sml b/src/list/src/numposrepScript.sml index 3446fa2dd1..2b5d8bcff3 100644 --- a/src/list/src/numposrepScript.sml +++ b/src/list/src/numposrepScript.sml @@ -422,4 +422,88 @@ val l2n_11 = Q.store_thm("l2n_11", \\ NTAC 2 (POP_ASSUM SUBST1_TAC) \\ FULL_SIMP_TAC (srw_ss()) [l2n_APPEND]]); +Theorem BITWISE_l2n_2: + LENGTH l1 = LENGTH l2 ==> + BITWISE (LENGTH l1) op (l2n 2 l1) (l2n 2 l2) = + l2n 2 (MAP2 (\x y. bool_to_bit $ op (ODD x) (ODD y)) l1 l2) +Proof + Q.ID_SPEC_TAC`l2` + \\ Induct_on`l1` + \\ simp[BITWISE_EVAL] + >- simp[BITWISE_def, l2n_def] + \\ Q.X_GEN_TAC`b` + \\ Cases \\ fs[BITWISE_EVAL] + \\ strip_tac + \\ fs[l2n_def] + \\ simp[SBIT_def, ODD_ADD, ODD_MULT, GSYM bool_to_bit_def] +QED + +Theorem l2n_2_neg: + !ls. + EVERY ($> 2) ls ==> + l2n 2 (MAP (\x. 1 - x) ls) = 2 ** LENGTH ls - SUC (l2n 2 ls) +Proof + Induct + \\ rw[l2n_def] + \\ fs[EXP, ADD1] + \\ simp[LEFT_SUB_DISTRIB, LEFT_ADD_DISTRIB, SUB_RIGHT_ADD] + \\ Q.SPECL_THEN[`ls`,`2`]mp_tac l2n_lt + \\ simp[] +QED + +Theorem l2n_max: + 0 < b ==> + !ls. (l2n b ls = b ** (LENGTH ls) - 1) <=> + (EVERY ((=) (b - 1) o flip $MOD b) ls) +Proof + strip_tac + \\ Induct + \\ rw[l2n_def] + \\ rw[EXP] + \\ Q.MATCH_GOALSUB_ABBREV_TAC`b * l + a` + \\ Q.MATCH_GOALSUB_ABBREV_TAC`b ** n` + \\ fs[EQ_IMP_THM] + \\ conj_tac + >- ( + strip_tac + \\ Cases_on`n=0` \\ fs[] >- (fs[Abbr`n`] \\ rw[]) + \\ `0 < b * b ** n` by simp[] + \\ `a + b * l + 1 = b * b ** n` by simp[] + \\ `(b * b ** n) DIV b = b ** n` by simp[MULT_TO_DIV] + \\ `(b * l + (a + 1)) DIV b = b ** n` by fs[] + \\ `(b * l) MOD b = 0` by simp[] + \\ `(b * l + (a + 1)) DIV b = (b * l) DIV b + (a + 1) DIV b` + by ( irule ADD_DIV_RWT \\ simp[] ) + \\ pop_assum SUBST_ALL_TAC + \\ `(a + 1) DIV b = if a = b - 1 then 1 else 0` + by ( + rw[] + \\ `a + 1 < b` suffices_by rw[DIV_EQ_0] + \\ simp[Abbr`a`] + \\ `h MOD b < b - 1` suffices_by simp[] + \\ `h MOD b < b` by simp[] + \\ fs[] ) + \\ `b * l DIV b = l` by simp[MULT_TO_DIV] + \\ pop_assum SUBST_ALL_TAC + \\ pop_assum SUBST_ALL_TAC + \\ `l < b ** n` by ( map_every Q.UNABBREV_TAC[`l`,`n`] + \\ irule l2n_lt \\ simp[] ) + \\ Cases_on`a = b - 1` \\ fs[] ) + \\ strip_tac + \\ rewrite_tac[LEFT_SUB_DISTRIB] + \\ Q.PAT_X_ASSUM`_ = a`(SUBST1_TAC o SYM) + \\ fs[SUB_LEFT_ADD] \\ rw[] + \\ Cases_on`n=0` \\ fs[] + \\ `b ** n <= b ** 0` by simp[] + \\ fs[EXP_BASE_LE_IFF] +QED + +Theorem l2n_PAD_RIGHT_0[simp]: + 0 < b ==> l2n b (PAD_RIGHT 0 h ls) = l2n b ls +Proof + Induct_on`ls` \\ rw[l2n_def, PAD_RIGHT, l2n_eq_0, EVERY_GENLIST] + \\ fs[PAD_RIGHT, l2n_APPEND] + \\ fs[l2n_eq_0, EVERY_GENLIST] +QED + val _ = export_theory() diff --git a/src/list/src/rich_listScript.sml b/src/list/src/rich_listScript.sml index 124bea3e0f..6a3de4d876 100644 --- a/src/list/src/rich_listScript.sml +++ b/src/list/src/rich_listScript.sml @@ -8,7 +8,7 @@ open HolKernel Parse boolLib BasicProvers; open numLib metisLib simpLib combinTheory arithmeticTheory prim_recTheory pred_setTheory listTheory pairTheory markerLib TotalDefn; -local open listSimps pred_setSimps in end; +local open listSimps pred_setSimps dep_rewrite in end; val FILTER_APPEND = FILTER_APPEND_DISTRIB val REVERSE = REVERSE_SNOC_DEF @@ -3189,6 +3189,201 @@ Proof \\ Cases_on`q` \\ fs[ADD1] QED +Theorem chunks_append_divides: + !n l1 l2. + 0 < n /\ divides n (LENGTH l1) /\ ~NULL l1 /\ ~NULL l2 ==> + chunks n (l1 ++ l2) = chunks n l1 ++ chunks n l2 +Proof + HO_MATCH_MP_TAC chunks_ind + \\ rw[dividesTheory.divides_def, PULL_EXISTS] + \\ simp[Once chunks_def] + \\ Cases_on`q=0` \\ fs[] \\ rfs[] + \\ IF_CASES_TAC + >- ( Cases_on`q` \\ fs[ADD1, LEFT_ADD_DISTRIB] \\ fs[LESS_OR_EQ] ) + \\ simp[DROP_APPEND, TAKE_APPEND] + \\ Q.MATCH_GOALSUB_ABBREV_TAC`lhs = _` + \\ simp[Once chunks_def] + \\ Cases_on`q = 1` \\ fs[] + >- ( + simp[Abbr`lhs`] + \\ fs[NOT_LESS_EQUAL] + \\ simp[DROP_LENGTH_TOO_LONG]) + \\ simp[Abbr`lhs`] + \\ `n - n * q = 0` by simp[] + \\ simp[] + \\ first_x_assum irule + \\ simp[NULL_EQ] + \\ qexists_tac`q - 1` + \\ simp[] +QED + +Theorem chunks_length[simp]: + chunks (LENGTH ls) ls = [ls] +Proof + rw[Once chunks_def] +QED + +Theorem chunks_not_nil[simp]: + !n ls. chunks n ls <> [] +Proof + HO_MATCH_MP_TAC chunks_ind + \\ rw[] + \\ rw[Once chunks_def] +QED + +Theorem LENGTH_chunks: + !n ls. 0 < n /\ ~NULL ls ==> + LENGTH (chunks n ls) = + LENGTH ls DIV n + (bool_to_bit $ ~divides n (LENGTH ls)) +Proof + HO_MATCH_MP_TAC chunks_ind + \\ rw[] + \\ rw[Once chunks_def, dividesTheory.DIV_EQUAL_0, bool_to_bit_def, + dividesTheory.divides_def] + \\ fs[LESS_OR_EQ, ADD1, NULL_EQ, bool_to_bit_def] \\ rfs[] + \\ rw[] + \\ fs[dividesTheory.divides_def, dividesTheory.SUB_DIV] + \\ rfs[] + >- ( + Cases_on`LENGTH ls DIV n = 0` >- rfs[dividesTheory.DIV_EQUAL_0] + \\ simp[] ) + >- ( + Cases_on`q` \\ fs[MULT_SUC] + \\ Q.MATCH_ASMSUB_RENAME_TAC`n + n * p` + \\ first_x_assum(Q.SPEC_THEN`2 + p`mp_tac) + \\ simp[LEFT_ADD_DISTRIB] ) + >- ( + first_x_assum(Q.SPEC_THEN`PRE q`mp_tac) + \\ Cases_on`q` \\ fs[MULT_SUC] ) + \\ Cases_on`q` \\ fs[MULT_SUC] + \\ simp[ADD_DIV_RWT] +QED + +Theorem EL_chunks: + !k ls n. + n < LENGTH (chunks k ls) /\ 0 < k /\ ~NULL ls ==> + EL n (chunks k ls) = TAKE k (DROP (n * k) ls) +Proof + HO_MATCH_MP_TAC chunks_ind \\ rw[NULL_EQ] + \\ Q.PAT_X_ASSUM`_ < LENGTH _ `mp_tac + \\ rw[Once chunks_def] \\ fs[] + \\ rw[Once chunks_def] + \\ Q.MATCH_GOALSUB_RENAME_TAC`EL m _` + \\ Cases_on`m` \\ fs[] + \\ pop_assum mp_tac + \\ dep_rewrite.DEP_REWRITE_TAC[LENGTH_chunks] + \\ simp[NULL_EQ] + \\ strip_tac + \\ dep_rewrite.DEP_REWRITE_TAC[DROP_DROP] + \\ simp[MULT_SUC] + \\ Q.MATCH_GOALSUB_RENAME_TAC`k + k * m <= _` + \\ `k * m <= LENGTH ls - k` suffices_by simp[] + \\ `m <= (LENGTH ls - k) DIV k` suffices_by simp[X_LE_DIV] + \\ fs[bool_to_bit_def] + \\ pop_assum mp_tac \\ rw[] +QED + +Theorem chunks_MAP: + !n ls. chunks n (MAP f ls) = MAP (MAP f) (chunks n ls) +Proof + HO_MATCH_MP_TAC chunks_ind \\ rw[] + \\ rw[Once chunks_def] + >- rw[Once chunks_def] + >- rw[Once chunks_def] + \\ fs[] + \\ simp[GSYM MAP_DROP] + \\ CONV_TAC(RAND_CONV(SIMP_CONV(srw_ss())[Once chunks_def])) + \\ simp[MAP_TAKE] +QED + +Theorem chunks_ZIP: + !n ls l2. LENGTH ls = LENGTH l2 ==> + chunks n (ZIP (ls, l2)) = MAP ZIP (ZIP (chunks n ls, chunks n l2)) +Proof + HO_MATCH_MP_TAC chunks_ind \\ rw[] + \\ rw[Once chunks_def] + >- ( rw[Once chunks_def] \\ rw[Once chunks_def] ) + >- rw[Once chunks_def] + \\ fs[] + \\ simp[GSYM ZIP_DROP] + \\ CONV_TAC(RAND_CONV(SIMP_CONV(srw_ss())[Once chunks_def])) + \\ CONV_TAC(PATH_CONV"rrrr"(SIMP_CONV(srw_ss())[Once chunks_def])) + \\ simp[ZIP_TAKE] +QED + +Theorem chunks_TAKE: + !n ls m. divides n m /\ 0 < m ==> + chunks n (TAKE m ls) = TAKE (m DIV n) (chunks n ls) +Proof + HO_MATCH_MP_TAC chunks_ind \\ rw[] + \\ CONV_TAC(RAND_CONV(SIMP_CONV(srw_ss())[Once chunks_def])) + \\ rw[] + >- ( + rw[Once chunks_def] \\ fs[LENGTH_TAKE_EQ] + \\ fs[dividesTheory.divides_def] + \\ BasicProvers.VAR_EQ_TAC + \\ Q.MATCH_GOALSUB_RENAME_TAC`n * m` + \\ fs[ZERO_LESS_MULT] + \\ `n <= n * m` by simp[LE_MULT_CANCEL_LBARE] + \\ dep_rewrite.DEP_REWRITE_TAC[TAKE_LENGTH_TOO_LONG] + \\ simp[MULT_DIV] ) + >- fs[dividesTheory.divides_def] + \\ fs[] + \\ simp[Once chunks_def, LENGTH_TAKE_EQ] + \\ `n <= m` by ( + rfs[dividesTheory.divides_def] \\ rw[] + \\ fs[ZERO_LESS_MULT] ) + \\ IF_CASES_TAC + >- ( + pop_assum mp_tac \\ rw[] + \\ `m = n` by fs[] \\ rw[] ) + \\ fs[TAKE_TAKE, DROP_TAKE] + \\ first_x_assum(Q.SPEC_THEN`m - n`mp_tac) + \\ simp[] + \\ impl_keep_tac >- ( + fs[dividesTheory.divides_def] + \\ qexists_tac`q - 1` + \\ simp[LEFT_SUB_DISTRIB] ) + \\ rw[] + \\ `m DIV n <> 0` by fs[dividesTheory.DIV_EQUAL_0] + \\ Cases_on`m DIV n` \\ fs[TAKE_TAKE_MIN] + \\ `MIN n m = n` by fs[MIN_DEF] \\ rw[] + \\ simp[dividesTheory.SUB_DIV] +QED + +Definition chunks_tr_aux_def: + chunks_tr_aux n ls acc = + if LENGTH ls <= SUC n then REVERSE $ ls :: acc + else chunks_tr_aux n (DROP (SUC n) ls) (TAKE (SUC n) ls :: acc) +Termination + WF_REL_TAC`measure $ LENGTH o FST o SND` + \\ rw[LENGTH_DROP] +End + +Definition chunks_tr_def: + chunks_tr n ls = if n = 0 then [ls] else chunks_tr_aux (n - 1) ls [] +End + +Theorem chunks_tr_aux_thm: + !n ls acc. + chunks_tr_aux n ls acc = + REVERSE acc ++ chunks (SUC n) ls +Proof + HO_MATCH_MP_TAC chunks_tr_aux_ind + \\ rw[] + \\ rw[Once chunks_tr_aux_def] + >- rw[Once chunks_def] + \\ CONV_TAC(RAND_CONV(SIMP_CONV(srw_ss())[Once chunks_def])) + \\ rw[] +QED + +Theorem chunks_tr_thm: + chunks_tr = chunks +Proof + simp[FUN_EQ_THM, chunks_tr_def] + \\ Cases \\ rw[chunks_tr_aux_thm] +QED + (*---------------------------------------------------------------------------*) (* Various lemmas from the CakeML project https://cakeml.org *) (*---------------------------------------------------------------------------*) @@ -3588,6 +3783,26 @@ Proof QED (* END more lemmas of IS_SUFFIX *) +Theorem IS_SUFFIX_dropWhile: + IS_SUFFIX ls (dropWhile P ls) +Proof + Induct_on`ls` + \\ rw[IS_SUFFIX_CONS] +QED + +Theorem LENGTH_dropWhile_id: + (LENGTH (dropWhile P ls) = LENGTH ls) <=> (dropWhile P ls = ls) +Proof + rw[EQ_IMP_THM] + \\ rw[dropWhile_id] + \\ Cases_on`ls` \\ fs[] + \\ strip_tac \\ fs[] + \\ `IS_SUFFIX t (dropWhile P t)` by simp[IS_SUFFIX_dropWhile] + \\ fs[IS_SUFFIX_APPEND] + \\ `LENGTH t = LENGTH l + LENGTH (dropWhile P t)` by metis_tac[LENGTH_APPEND] + \\ fs[] +QED + Theorem nub_GENLIST: nub (GENLIST f n) = MAP f (FILTER (\i. !j. (i < j) /\ (j < n) ==> f i <> f j) (COUNT_LIST n)) diff --git a/src/monad/more_monads/state_transformerScript.sml b/src/monad/more_monads/state_transformerScript.sml index 8dcaec9ab7..55fada8b09 100644 --- a/src/monad/more_monads/state_transformerScript.sml +++ b/src/monad/more_monads/state_transformerScript.sml @@ -2,9 +2,6 @@ open HolKernel Parse boolLib pairTheory pairSyntax combinTheory listTheory; val _ = new_theory "state_transformer" -infixr 0 || -infix 1 >>; - val DEF = Lib.with_flag (boolLib.def_suffix, "_DEF") TotalDefn.Define (* ------------------------------------------------------------------------- *) diff --git a/src/n-bit/wordsScript.sml b/src/n-bit/wordsScript.sml index 7d17c753b9..2389480103 100644 --- a/src/n-bit/wordsScript.sml +++ b/src/n-bit/wordsScript.sml @@ -5060,6 +5060,224 @@ QED Theorem CARD_WORD = CARD_CART_UNIV |> INST_TYPE [alpha |-> bool] |> SIMP_RULE bool_ss [CARD_BOOL,FINITE_BOOL] +Theorem MEM_w2l_less: + 1 < b /\ MEM x (w2l b w) ==> x < b +Proof + Cases_on`w` + \\ rw[w2l_def] + \\ rpt $ pop_assum mp_tac + \\ map_every qid_spec_tac [`x`,`n`,`b`] + \\ recInduct n2l_ind + \\ rw[] + \\ pop_assum mp_tac + \\ rw[Once n2l_def] + >- ( irule MOD_LESS \\ gs[] ) + \\ first_x_assum irule + \\ gs[DIV_LT_X] + \\ Cases_on`b` \\ gs[MULT] +QED + +Theorem l2w_PAD_RIGHT_0[simp]: + 0 < b ==> l2w b (PAD_RIGHT 0 h ls) = l2w b ls +Proof + rw[l2w_def] +QED + +Theorem word_and_lsl_eq_0: + w2n w1 < 2 ** n ==> w1 && w2 << n = 0w +Proof + Cases_on`w1` \\ Cases_on`w2` + \\ rw[word_and_n2w, word_lsl_n2w] + \\ drule BITWISE_AND_SHIFT_EQ_0 + \\ simp[] +QED + +(* ------------------------------------------------------------------------- + Theorems about word_{to,from}_bin_list + ------------------------------------------------------------------------- *) + +Theorem LENGTH_word_to_bin_list_bound: + LENGTH (word_to_bin_list (w:'a word)) <= (dimindex (:'a)) +Proof + rw[word_to_bin_list_def, w2l_def, LENGTH_n2l] + \\ Cases_on`w` \\ simp[] + \\ fs[dimword_def, GSYM LESS_EQ, LOG2_def, logrootTheory.LT_EXP_LOG] +QED + +Theorem word_from_bin_list_ror: + x < dimindex(:'a) /\ LENGTH ls = dimindex(:'a) + ==> + word_ror (word_from_bin_list ls : 'a word) x = + word_from_bin_list (DROP x ls ++ TAKE x ls) +Proof + rw[word_from_bin_list_def, l2w_def] + \\ Cases_on`x = 0` \\ gs[SHIFT_ZERO] + \\ rw[word_ror_n2w, l2n_APPEND] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ qspecl_then[`x`,`ls`](mp_tac o SYM) listTheory.TAKE_DROP + \\ disch_then(SUBST1_TAC o Q.AP_TERM`l2n 2`) + \\ rewrite_tac[l2n_APPEND] + \\ simp[] + \\ qmatch_goalsub_abbrev_tac`BITS _ _ (2 ** x * ld + lt)` + \\ qspecl_then[`x - 1`,`0`,`ld`,`lt`]mp_tac BITS_SUM2 + \\ simp[ADD1] + \\ disch_then kall_tac + \\ qspecl_then[`x-1`,`lt`]mp_tac BITS_ZEROL + \\ simp[ADD1] + \\ impl_keep_tac + >- ( + qspecl_then[`TAKE x ls`,`2`]mp_tac l2n_lt + \\ simp[] ) + \\ simp[] + \\ disch_then kall_tac + \\ simp_tac std_ss [Once ADD_COMM, SimpRHS] + \\ qmatch_goalsub_abbrev_tac`BITS h x` + \\ qspecl_then[`h`,`x`,`ld`,`lt`]mp_tac BITS_SUM + \\ simp[] \\ disch_then kall_tac + \\ DEP_REWRITE_TAC[BITS_ZERO4] + \\ simp[Abbr`h`] + \\ DEP_REWRITE_TAC[BITS_ZEROL] + \\ qspecl_then[`DROP x ls`,`2`]mp_tac l2n_lt + \\ simp[ADD1] +QED + +Theorem word_from_bin_list_rol: + x < dimindex(:'a) /\ LENGTH ls = dimindex(:'a) + ==> + word_rol (word_from_bin_list ls : 'a word) x = + word_from_bin_list (LASTN x ls ++ BUTLASTN x ls) +Proof + rw[word_rol_def] + \\ Cases_on`x=0` + >- ( + rw[rich_listTheory.LASTN, rich_listTheory.BUTLASTN] + \\ ONCE_REWRITE_TAC[GSYM ROR_MOD] + \\ rw[SHIFT_ZERO] ) + \\ DEP_REWRITE_TAC[word_from_bin_list_ror] + \\ simp[rich_listTheory.LASTN_DROP, rich_listTheory.BUTLASTN_TAKE] +QED + +Theorem word_from_bin_list_and: + LENGTH l1 = dimindex(:'a) /\ + LENGTH l2 = dimindex(:'a) + ==> + word_from_bin_list l1 && word_from_bin_list l2 : 'a word = + word_from_bin_list (MAP2 (\x y. bool_to_bit $ (ODD x /\ ODD y)) l1 l2) +Proof + rw[word_from_bin_list_def, l2w_def, word_and_n2w] + \\ qmatch_goalsub_abbrev_tac`BITWISE n` + \\ qmatch_goalsub_abbrev_tac`a MOD d = b MOD d` + \\ `d = 2 ** n` + by simp[Abbr`d`, Abbr`n`, dimword_def] + \\ `a < d` by ( + pop_assum SUBST1_TAC + \\ qunabbrev_tac`a` + \\ irule BITWISE_LT_2EXP ) + \\ `b < d` + by ( + qunabbrev_tac`b` + \\ qmatch_goalsub_abbrev_tac`l2n 2 ls` + \\ `n = LENGTH ls` by simp[Abbr`ls`] + \\ qunabbrev_tac`d` + \\ qpat_x_assum`_ = 2 ** _`SUBST1_TAC + \\ pop_assum SUBST1_TAC + \\ irule l2n_lt \\ simp[] ) + \\ simp[] + \\ unabbrev_all_tac + \\ DEP_REWRITE_TAC[GSYM BITWISE_l2n_2] + \\ simp[] +QED + +Theorem word_from_bin_list_not: + LENGTH ls = dimindex(:'a) /\ EVERY ($> 2) ls ==> + ~word_from_bin_list ls : 'a word = + word_from_bin_list (MAP (\x. 1 - x) ls) +Proof + rw[word_from_bin_list_def, l2w_def] + \\ rewrite_tac[word_1comp_n2w] + \\ Cases_on`l2n 2 ls = dimword(:'a) - 1` + >- ( + mp_then Any (qspec_then`ls`mp_tac) + l2n_max (SIMP_CONV(srw_ss())[]``0 < 2n`` |> EQT_ELIM) + \\ `dimword (:'a) = 2 ** LENGTH ls` by simp[dimword_def] + \\ first_assum (SUBST_ALL_TAC o SYM) + \\ pop_assum kall_tac + \\ pop_assum SUBST_ALL_TAC + \\ simp[ADD1, ZERO_LT_dimword] + \\ strip_tac + \\ qmatch_goalsub_abbrev_tac`A MOD N = 0` + \\ `A = 0` suffices_by rw[] + \\ rw[Abbr`A`, l2n_eq_0] + \\ gs[listTheory.EVERY_MAP, listTheory.EVERY_MEM] + \\ rw[] + \\ `2 > x` by simp[] + \\ `x < 2` by simp[] + \\ pop_assum mp_tac + \\ `x MOD 2 = 1` by simp[] + \\ rewrite_tac[NUMERAL_LESS_THM] + \\ strip_tac \\ fs[] ) + \\ `SUC (l2n 2 ls) < dimword(:'a)` + by ( + qspecl_then[`ls`,`2`]mp_tac l2n_lt + \\ gs[dimword_def] ) + \\ qmatch_goalsub_abbrev_tac`_ = n2w $ l2n 2 l1` + \\ `l2n 2 l1 < dimword(:'a)` + by ( + qspecl_then[`l1`,`2`]mp_tac l2n_lt + \\ gs[dimword_def, Abbr`l1`] ) + \\ simp[Abbr`l1`] + \\ drule l2n_2_neg + \\ simp[dimword_def] +QED + +Theorem word_from_bin_list_xor: + LENGTH b1 = LENGTH b2 ==> + word_from_bin_list b1 ?? word_from_bin_list b2 = + word_from_bin_list (MAP (\(x,y). (x + y) MOD 2) (ZIP (b1, b2))) +Proof + qid_spec_tac`b2` + \\ Induct_on`b1` + \\ rw[] + >- (EVAL_TAC \\ rw[WORD_XOR_CLAUSES]) + \\ Cases_on`b2` \\ gs[] + \\ gs[word_from_bin_list_def, l2w_def, l2n_def] + \\ gs[GSYM word_add_n2w, GSYM word_mul_n2w] + \\ first_x_assum(qspec_then`t`(mp_tac o GSYM)) + \\ simp[] \\ disch_then kall_tac + \\ DEP_REWRITE_TAC[WORD_ADD_XOR] + \\ `n2w 2 = n2w (2 ** 1)` by simp[] + \\ pop_assum SUBST_ALL_TAC + \\ simp[GSYM WORD_MUL_LSL] + \\ rewrite_tac[LSL_BITWISE] + \\ DEP_REWRITE_TAC[word_and_lsl_eq_0] + \\ simp[] + \\ conj_tac + >- ( + rw[] + \\ Cases_on`2 < dimword(:'a)` \\ gs[MOD_LESS, MOD_MOD_LESS_EQ] + \\ Cases_on`dimword(:'a) = 2` \\ gs[MOD_MOD] + \\ `dimword(:'a) = 1` by simp[ZERO_LT_dimword] + \\ gs[] ) + \\ rewrite_tac[GSYM WORD_XOR_ASSOC, GSYM LSL_BITWISE] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ rewrite_tac[Once WORD_XOR_COMM] + \\ rewrite_tac[GSYM WORD_XOR_ASSOC] + \\ AP_THM_TAC \\ AP_TERM_TAC + \\ qmatch_goalsub_rename_tac`x + y` + \\ `x MOD 2 < 2 /\ y MOD 2 < 2` by simp[] + \\ ntac 2 (pop_assum mp_tac) + \\ rewrite_tac[NUMERAL_LESS_THM] + \\ Cases_on`ODD (x + y)` + >- ( + `(x + y) MOD 2 = 1` by gs[ODD_MOD2_LEM] + \\ gs[ODD_ADD] + \\ Cases_on`ODD x` \\ gs[ODD_MOD2_LEM, WORD_XOR_CLAUSES] ) + \\ gs[GSYM EVEN_ODD] + \\ drule (iffLR EVEN_ADD) + \\ gs[EVEN_MOD2] + \\ Cases_on`x MOD 2 = 0` \\ gs[WORD_XOR_CLAUSES] +QED + (* ------------------------------------------------------------------------- Create a few word sizes ------------------------------------------------------------------------- *) diff --git a/src/num/extra_theories/bitScript.sml b/src/num/extra_theories/bitScript.sml index 8ff867a02a..3302624813 100644 --- a/src/num/extra_theories/bitScript.sml +++ b/src/num/extra_theories/bitScript.sml @@ -381,6 +381,12 @@ val NOT_MOD2_LEM2 = Q.store_thm("NOT_MOD2_LEM2", val ODD_MOD2_LEM = Q.store_thm("ODD_MOD2_LEM", `!n. ODD n = ((n MOD 2) = 1)`, RW_TAC arith_ss [ODD_EVEN, MOD_2]) +Theorem ODD_MOD_2[simp]: + ODD (x MOD 2) = ODD x +Proof + RW_TAC arith_ss [ODD_MOD2_LEM] +QED + (* ------------------------------------------------------------------------- *) val DIV_MULT_THM = Q.store_thm("DIV_MULT_THM", @@ -1520,4 +1526,55 @@ Proof simp_tac std_ss [] QED +Theorem BITWISE_COMM: + (!m. m <= n ==> op (BIT m x) (BIT m y) = op (BIT m y) (BIT m x)) + ==> BITWISE n op x y = BITWISE n op y x +Proof + Induct_on`n` + \\ SRW_TAC[][BITWISE_def] + \\ first_assum(Q.SPEC_THEN`n`mp_tac) + \\ impl_tac >- SRW_TAC[][] + \\ disch_then SUBST1_TAC + \\ simp[] + \\ first_x_assum irule + \\ SRW_TAC[][] + \\ first_x_assum irule + \\ simp[] +QED + +Triviality BITWISE_AND_0_lemma: + BITWISE w $/\ x 0 = 0 +Proof + Q.ID_SPEC_TAC`x` + \\ Induct_on`w` + \\ SRW_TAC[][BITWISE_def, SBIT_def, BIT_ZERO] +QED + +Theorem BITWISE_AND_0[simp]: + BITWISE w $/\ x 0 = 0 /\ + BITWISE w $/\ 0 x = 0 +Proof + Q.SPECL_THEN[`$/\`,`0`,`x`,`w`]mp_tac(Q.GENL[`op`,`x`,`y`,`n`]BITWISE_COMM) + \\ impl_tac >- SRW_TAC[][BIT_ZERO] + \\ disch_then SUBST1_TAC + \\ SRW_TAC[][BITWISE_AND_0_lemma] +QED + +Theorem BITWISE_AND_SHIFT_EQ_0: + !w x y n. + x < 2 ** n ==> + BITWISE w $/\ x (y * 2 ** n) = 0 +Proof + Induct \\ SRW_TAC[][BITWISE_def, SBIT_def] + \\ strip_tac + \\ Cases_on`w < n` + >- ( drule BIT_SHIFT_THM3 \\ simp[] + \\ Q.EXISTS_TAC`y` \\ simp[]) + \\ FULL_SIMP_TAC(srw_ss())[NOT_LESS] + \\ drule TWOEXP_MONO2 \\ strip_tac + \\ `x < 2 ** w` by METIS_TAC[LESS_LESS_EQ_TRANS] + \\ drule NOT_BIT_GT_TWOEXP + \\ simp[] +QED + val _ = export_theory() diff --git a/src/num/termination/TotalDefn.sml b/src/num/termination/TotalDefn.sml index c1813d97f1..2b47c6312b 100644 --- a/src/num/termination/TotalDefn.sml +++ b/src/num/termination/TotalDefn.sml @@ -738,16 +738,17 @@ val tDefine = located_tDefine DB.Unknown version of Define that allows control of options with "attributes" attached to the name, and provides an optional termination tactic ---------------------------------------------------------------------- *) -fun test_remove n sl = - if mem n sl then (true, set_diff sl [n]) else (false, sl) +fun test_remove n kvl = + let val (ns, rest) = Lib.partition (fn (k,_) => k = n) kvl + in + (not (null ns), rest) + end -fun find_indoption sl = - case List.find (String.isPrefix "induction=") sl of - NONE => (NONE, sl) - | SOME s => ( - SOME (String.extract(s,size "induction=",NONE)), - set_diff sl [s] - ) +fun find_indoption kvl = + case Lib.partition (fn (k,_) => k = "induction") kvl of + ([], _) => (NONE, kvl) + | ([(k,[v])], rest) => (SOME v, rest) + | _ => raise ERR "Definition" "multiple induction attribute-values" fun tailrecDefine loc nm q = let @@ -757,16 +758,22 @@ fun tailrecDefine loc nm q = Defn.add_defs_to_EVAL [(nm,th)]; th end - +val _ = List.app ThmAttribute.reserve_word + ["nocompute", "schematic", "tailrecursive"] fun located_qDefine loc stem q tacopt = let - val (corename, attrs) = ThmAttribute.extract_attributes stem - val (nocomp, attrs) = test_remove "nocompute" attrs - val (svarsok, attrs) = test_remove "schematic" attrs - val (notuserdef, attrs) = test_remove "notuserdef" attrs - val (rebindok, attrs) = test_remove "allow_rebind" attrs - val (tailrecp, attrs) = test_remove "tailrecursive" attrs - val (indopt,attrs) = find_indoption attrs + val {thmname=corename, attrs=attrs,reserved=R,unknown} = + ThmAttribute.extract_attributes stem + val (nocomp, R) = test_remove "nocompute" R + val (svarsok, R) = test_remove "schematic" R + val (notuserdef, R) = test_remove "notuserdef" R + val (rebindok, R) = test_remove "allow_rebind" R + val (tailrecp, R) = test_remove "tailrecursive" R + val (indopt, R) = find_indoption R + val _ = null unknown orelse + raise ERR "Definition" + ("Unknown attributes: " ^ + String.concatWith ", " (map #1 unknown)) fun fmod f = f |> (if nocomp then trace ("computeLib.auto_import_definitions", 0) else (fn f => f)) @@ -786,10 +793,10 @@ fun located_qDefine loc stem q tacopt = \no sense" | (false, NONE) => fmod (located_xDefine loc corename) q | (false, SOME tac) => fmod (located_tDefine loc corename q) tac - fun proc_attr a = - ThmAttribute.store_at_attribute{name = corename, attrname = a, - thm = thm} - val attrs = if notuserdef then attrs else "userdef" :: attrs + fun proc_attr (k,vs) = + ThmAttribute.store_at_attribute{name = corename, attrname = k, + args = vs, thm = thm} + val attrs = if notuserdef then attrs else ("userdef",[]) :: attrs val gen_ind = if tailrecp then (fn th => raise ERR "Unseen" "") else Prim_rec.gen_indthm {lookup_ind = TypeBase.induction_of} diff --git a/src/num/theories/arithmeticScript.sml b/src/num/theories/arithmeticScript.sml index 9ca58af265..5d7c49a333 100644 --- a/src/num/theories/arithmeticScript.sml +++ b/src/num/theories/arithmeticScript.sml @@ -224,6 +224,11 @@ val NRC = new_recursive_definition { def = “(NRC R 0 x y = (x = y)) /\ (NRC R (SUC n) x y = ?z. R x z /\ NRC R n z y)”}; +val bool_to_bit_def = new_definition( + "bool_to_bit_def", + “bool_to_bit b = if b then 1 else 0” +); + val _ = overload_on ("RELPOW", “NRC”) val _ = overload_on ("NRC", “NRC”) @@ -2962,6 +2967,19 @@ Proof >> PROVE_TAC [] QED +Theorem MOD_DIV_SAME[simp]: + 0 < y ==> x MOD y DIV y = 0 +Proof + strip_tac + \\ Cases_on`y=1` \\ fs[DIV_1, MOD_1] + \\ `1 < y` by ( + simp[ONE, SUC_LT] + \\ Cases_on`y` \\ fs[NOT_LESS_0, ONE] + \\ drule (NOT_LT_ZERO_EQ_ZERO |> GSYM |> SPEC_ALL |> iffRL |> CONTRAPOS) + \\ rw[] ) + \\ rw[DIV_EQ_0, MOD_LESS] +QED + (* ------------------------------------------------------------------------ *) (* Some miscellaneous lemmas (from transc.ml) *) (* ------------------------------------------------------------------------ *) @@ -4997,4 +5015,34 @@ val MOD_EXP = store_thm( `_ = (a * a ** m) MOD n` by rw[MOD_TIMES2] >> rw[EXP]); +Theorem ODD_bool_to_bit[simp]: + ODD (bool_to_bit b) = b /\ + ODD (1 - bool_to_bit b) = ~b +Proof + rw[bool_to_bit_def, ODD, ONE, SUB_MONO_EQ, SUB_0] +QED + +Theorem bool_to_bit_neq_add: + bool_to_bit (x <> y) = + (bool_to_bit x + bool_to_bit y) MOD 2 +Proof + reverse(rw[bool_to_bit_def, ADD_0, ADD]) + >- ( + rw[ONE, ADD] \\ rw[GSYM TWO, GSYM ONE] + \\ irule EQ_SYM + \\ irule (DIVMOD_ID |> SPEC_ALL |> UNDISCH |> cj 2 |> DISCH_ALL) + \\ rw[TWO, LESS_0] ) + \\ irule EQ_SYM + \\ irule ONE_MOD + \\ rw[ONE, TWO, LESS_MONO, LESS_0] +QED + +Theorem bool_to_bit_MOD_2[simp]: + bool_to_bit x MOD 2 = bool_to_bit x +Proof + rw[bool_to_bit_def] + \\ irule ONE_MOD + \\ rw[ONE, TWO, LESS_MONO, LESS_0] +QED + val _ = export_theory() diff --git a/src/num/theories/cv_compute/automation/cv_primScript.sml b/src/num/theories/cv_compute/automation/cv_primScript.sml index 3df66869bc..f6da846e3b 100644 --- a/src/num/theories/cv_compute/automation/cv_primScript.sml +++ b/src/num/theories/cv_compute/automation/cv_primScript.sml @@ -904,10 +904,16 @@ QED Theorem cv_rep_word_lsr[cv_rep]: from_word (word_lsr (w:'a word) n) = - cv_div (from_word w) (cv_exp (Num 2) (Num n)) + let k = Num n in + cv_if (cv_lt k (Num (dimindex (:'a)))) + (cv_div (from_word w) (cv_exp (Num 2) k)) + (Num 0) Proof gvs [cv_rep_def] \\ rw [] \\ gvs [] - \\ gvs [from_word_def,cv_exp_def,w2n_lsr] + >- gvs [from_word_def,cv_exp_def,w2n_lsr] + \\ gvs[from_word_def] + \\ irule LSR_LIMIT + \\ pop_assum mp_tac \\ rw[] QED Theorem word_asr_add[cv_inline]: @@ -1239,9 +1245,6 @@ Proof \\ completeInduct_on ‘m’ \\ simp [Once cv_word_and_loop_def] \\ Cases_on ‘m = 0’ \\ fs [wordsTheory.WORD_AND_CLAUSES] - >- - (qsuff_tac ‘!l n. BITWISE l $/\ 0 n = 0’ >- fs [] - \\ Induct \\ fs [BITWISE]) \\ gvs [PULL_FORALL,AND_IMP_INTRO] \\ rw [] \\ Cases_on ‘l’ \\ gvs [] \\ rename [‘n < 2 ** SUC l’] diff --git a/src/num/theories/cv_compute/automation/cv_typeScript.sml b/src/num/theories/cv_compute/automation/cv_typeScript.sml index a342a3b504..d2fb5a2936 100644 --- a/src/num/theories/cv_compute/automation/cv_typeScript.sml +++ b/src/num/theories/cv_compute/automation/cv_typeScript.sml @@ -190,7 +190,7 @@ Definition from_list_def: from_list f (x::xs) = Pair (f x) (from_list f xs) End -Definition to_list_def: +Definition to_list_def[nocompute]: to_list f (Num n) = [] /\ to_list f (Pair x y) = f x :: to_list f y End @@ -203,6 +203,23 @@ Proof \\ Induct \\ fs [from_list_def,to_list_def] QED +Definition to_list_tr_def: + to_list_tr f (Num n) acc = REVERSE acc /\ + to_list_tr f (Pair x y) acc = to_list_tr f y (f x :: acc) +End + +Theorem to_list_tr_eq_lemma[local]: + !v acc. to_list_tr f v acc = REVERSE acc ++ to_list f v +Proof + Induct \\ rw[to_list_def, to_list_tr_def] +QED + +Theorem to_list_tr_eq[compute]: + to_list f v = to_list_tr f v [] +Proof + rw[to_list_tr_eq_lemma] +QED + (* used in definitions of to-functions of user-defined datatype *) Definition cv_has_shape_def: diff --git a/src/parallel_builds/core/Holmakefile b/src/parallel_builds/core/Holmakefile index 4ef8ca1c3f..745e46a81e 100644 --- a/src/parallel_builds/core/Holmakefile +++ b/src/parallel_builds/core/Holmakefile @@ -32,7 +32,8 @@ ifdef HOLSELFTESTLEVEL # example directories to build at selftest level 1 EXDIRS = algebra/aat \ arm/arm6-verification arm/armv8-memory-model arm/experimental \ - CCS Crypto/RSA Crypto/SHA-2 Hoare-for-divergence MLsyntax \ + CCS Crypto/RSA Crypto/SHA-2 Crypto/pedersenCommitment Crypto/sigmaProtocol \ + Crypto/Keccak Hoare-for-divergence MLsyntax \ PSL/1.01/executable-semantics PSL/1.1/official-semantics \ STE algorithms computability countchars dependability dev \ developers/ThmSetData \ diff --git a/src/parse/Preterm.sml b/src/parse/Preterm.sml index 8a2e876051..531ed153a9 100644 --- a/src/parse/Preterm.sml +++ b/src/parse/Preterm.sml @@ -4,6 +4,8 @@ struct open Feedback Lib GrammarSpecials; open errormonad typecheck_error +infix >> >- + val ERR = mk_HOL_ERR "Preterm" val ERRloc = mk_HOL_ERRloc "Preterm" diff --git a/src/parse/term_pp.sml b/src/parse/term_pp.sml index 5c54f99f5f..4ddaee56d2 100644 --- a/src/parse/term_pp.sml +++ b/src/parse/term_pp.sml @@ -5,6 +5,8 @@ open Portable HolKernel term_grammar HOLtokens HOLgrammars GrammarSpecials PrecAnalysis +infix >> >- + val PP_ERR = mk_HOL_ERR "term_pp"; fun PRINT s = print (s ^ "\n") diff --git a/src/portableML/HOLsexp.sig b/src/portableML/HOLsexp.sig index df245dac1e..8f308d63ed 100644 --- a/src/portableML/HOLsexp.sig +++ b/src/portableML/HOLsexp.sig @@ -30,6 +30,8 @@ sig ('a * 'b * 'c * 'd) decoder val list_decode : 'a decoder -> 'a list decoder val tagged_decode : string -> 'a decoder -> 'a decoder + val map_decode : ('a -> 'b) -> 'a decoder -> 'b decoder + val bind_decode : 'a decoder -> ('a -> 'b option) -> 'b decoder val printer : t HOLPP.pprinter diff --git a/src/portableML/HOLsexp.sml b/src/portableML/HOLsexp.sml index fc4f3d884a..de98ebee28 100644 --- a/src/portableML/HOLsexp.sml +++ b/src/portableML/HOLsexp.sml @@ -17,6 +17,11 @@ struct val fromFile = HOLsexp_parser.raw_read_file val fromStream = HOLsexp_parser.raw_read_stream + fun map_decode f d = Option.map f o d + fun bind_decode d f x = + case d x of + NONE => NONE + | SOME y => f y fun r3_to_p12 (x,y,z) = (x, (y, z)) fun p12_to_r3 (x,(y,z)) = (x,y,z) fun pair_encode (a,b) (x,y) = Cons(a x, b y) @@ -32,7 +37,7 @@ struct Cons(Symbol s', rest) => if s = s' then SOME rest else NONE | _ => NONE - fun tagged_decode s dec t = Option.mapPartial dec (dest_tagged s t) + fun tagged_decode s dec = bind_decode (dest_tagged s) dec fun list_encode enc els = List (map enc els) fun break_list s = @@ -70,13 +75,12 @@ struct (fn (u,v,w,x) => (u,(v,(w,x)))) fun singleton [a] = SOME a | singleton _ = NONE - fun sing_decode d = - Option.mapPartial singleton o list_decode d + fun sing_decode d = bind_decode (list_decode d) singleton fun pair3_decode (a,b,c) = - Option.map p12_to_r3 o pair_decode(a, pair_decode(b,sing_decode c)) + map_decode p12_to_r3 (pair_decode(a, pair_decode(b,sing_decode c))) fun pair4_decode (a,b,c,d) = - Option.map (fn (u,(v,(w,x))) => (u,v,w,x)) o - pair_decode(a,pair_decode(b,pair_decode(c,sing_decode d))) + map_decode (fn (u,(v,(w,x))) => (u,v,w,x)) + (pair_decode(a,pair_decode(b,pair_decode(c,sing_decode d)))) fun is_list s = case s of diff --git a/src/portableML/Portable.sig b/src/portableML/Portable.sig index deed626de8..b2b7a19544 100644 --- a/src/portableML/Portable.sig +++ b/src/portableML/Portable.sig @@ -204,6 +204,7 @@ sig val ordof: string * int -> int val replace_string : {from:string,to:string} -> string -> string val remove_wspace : string -> string + val remove_external_wspace : string -> string val time_eq: time -> time -> bool val timestamp: unit -> time diff --git a/src/portableML/Portable.sml b/src/portableML/Portable.sml index 030e01ee87..5eeb8075df 100644 --- a/src/portableML/Portable.sml +++ b/src/portableML/Portable.sml @@ -603,6 +603,12 @@ fun replace_string {from, to} = val remove_wspace = String.translate (fn c => if Char.isSpace c then "" else str c) +fun remove_external_wspace s = + let + open Substring + in + string (dropl Char.isSpace (dropr Char.isSpace (full s))) + end (*--------------------------------------------------------------------------- System diff --git a/src/postkernel/TheoryReader.sml b/src/postkernel/TheoryReader.sml index 06576b48ba..71499f138d 100644 --- a/src/postkernel/TheoryReader.sml +++ b/src/postkernel/TheoryReader.sml @@ -41,8 +41,9 @@ val dep_decode = let | (n,[i]) :: rest => SOME{head = (n,i), deps = rest} | _ => NONE in - Option.mapPartial depmunge o - list_decode (pair_decode(string_decode, list_decode int_decode)) + bind_decode ( + list_decode (pair_decode(string_decode, list_decode int_decode)) + ) depmunge end val deptag_decode = let open HOLsexp in pair_decode(dep_decode, list_decode string_decode) @@ -53,8 +54,9 @@ val thm_decode = fun thmmunge(s,(di,tags),tms) = {name = s, depinfo = di, tagnames = tags, encoded_hypscon = tms} in - Option.map thmmunge o - pair3_decode (string_decode, deptag_decode, list_decode string_decode) + map_decode thmmunge ( + pair3_decode (string_decode, deptag_decode, list_decode string_decode) + ) end exception TheoryReader of string @@ -101,8 +103,9 @@ fun string_to_class "A" = SOME DB.Axm | string_to_class _ = NONE fun class_decode c = - Option.map (List.map (fn i => (i, c))) o - HOLsexp.list_decode HOLsexp.int_decode + HOLsexp.map_decode (List.map (fn i => (i, c))) ( + HOLsexp.list_decode HOLsexp.int_decode + ) fun load_thydata thyname path = let diff --git a/src/postkernel/ThyDataSexp.sig b/src/postkernel/ThyDataSexp.sig index ddd24b83aa..4693df49ba 100644 --- a/src/postkernel/ThyDataSexp.sig +++ b/src/postkernel/ThyDataSexp.sig @@ -71,6 +71,7 @@ val tag_encode : string -> ('a -> t) -> ('a -> t) val tag_decode : string -> 'a dec -> 'a dec val || : 'a dec * 'a dec -> 'a dec +val >> : 'a dec * ('a -> 'b) -> 'b dec val first : 'a dec list -> 'a dec end diff --git a/src/quotient/examples/finite_setScript.sml b/src/quotient/examples/finite_setScript.sml index 3637698764..c53e7f1c09 100644 --- a/src/quotient/examples/finite_setScript.sml +++ b/src/quotient/examples/finite_setScript.sml @@ -1,7 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* ------------------------------------------------------------------------ *) (* Representing finite sets as a new datatype in the HOL logic. *) diff --git a/src/quotient/examples/lambda/betaScript.sml b/src/quotient/examples/lambda/betaScript.sml index 8d1f3f02d4..9ef922bcdb 100644 --- a/src/quotient/examples/lambda/betaScript.sml +++ b/src/quotient/examples/lambda/betaScript.sml @@ -1,8 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - - val _ = new_theory "beta"; diff --git a/src/quotient/examples/lambda/etaScript.sml b/src/quotient/examples/lambda/etaScript.sml index a6217ac42c..7851869a45 100644 --- a/src/quotient/examples/lambda/etaScript.sml +++ b/src/quotient/examples/lambda/etaScript.sml @@ -1,8 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - - val _ = new_theory "eta"; diff --git a/src/quotient/examples/lambda/reductionScript.sml b/src/quotient/examples/lambda/reductionScript.sml index 87e55c47bd..5df89ab3d0 100644 --- a/src/quotient/examples/lambda/reductionScript.sml +++ b/src/quotient/examples/lambda/reductionScript.sml @@ -1,8 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - - val _ = new_theory "reduction"; diff --git a/src/quotient/examples/lambda/termScript.sml b/src/quotient/examples/lambda/termScript.sml index 67296b1980..2e2b6339f7 100644 --- a/src/quotient/examples/lambda/termScript.sml +++ b/src/quotient/examples/lambda/termScript.sml @@ -1,6 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; (* --------------------------------------------------------------------- *) diff --git a/src/quotient/examples/lambda/variableScript.sml b/src/quotient/examples/lambda/variableScript.sml index d4a07b716c..7b4e558f0b 100644 --- a/src/quotient/examples/lambda/variableScript.sml +++ b/src/quotient/examples/lambda/variableScript.sml @@ -15,9 +15,6 @@ (* Boilerplate. *) (* --------------------------------------------------------------------- *) open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* --------------------------------------------------------------------- *) (* Create the theory. *) diff --git a/src/quotient/examples/more_listScript.sml b/src/quotient/examples/more_listScript.sml index 5c73f2f7b7..92dd2c8a78 100644 --- a/src/quotient/examples/more_listScript.sml +++ b/src/quotient/examples/more_listScript.sml @@ -1,7 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - val _ = new_theory "more_list"; diff --git a/src/quotient/examples/sigma/alphaScript.sml b/src/quotient/examples/sigma/alphaScript.sml index df56757436..652d2736e9 100644 --- a/src/quotient/examples/sigma/alphaScript.sml +++ b/src/quotient/examples/sigma/alphaScript.sml @@ -1,7 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* --------------------------------------------------------------------- *) (* Building the definitions of alpha equivalence for object expressions. *) diff --git a/src/quotient/examples/sigma/objectScript.sml b/src/quotient/examples/sigma/objectScript.sml index f5d05d4a3e..6d4e5db795 100644 --- a/src/quotient/examples/sigma/objectScript.sml +++ b/src/quotient/examples/sigma/objectScript.sml @@ -1,7 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* --------------------------------------------------------------------- *) (* Embedding objects as a foundational layer, according to *) diff --git a/src/quotient/examples/sigma/reductionScript.sml b/src/quotient/examples/sigma/reductionScript.sml index 3228191038..eb04201cc0 100644 --- a/src/quotient/examples/sigma/reductionScript.sml +++ b/src/quotient/examples/sigma/reductionScript.sml @@ -1,7 +1,4 @@ open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* --------------------------------------------------------------------- *) (* Embedding the reduction semantaics of objects as a foundational *) diff --git a/src/quotient/examples/sigma/variableScript.sml b/src/quotient/examples/sigma/variableScript.sml index 77ca558333..c4aacea0ad 100644 --- a/src/quotient/examples/sigma/variableScript.sml +++ b/src/quotient/examples/sigma/variableScript.sml @@ -21,9 +21,6 @@ (* Boilerplate. *) (* --------------------------------------------------------------------- *) open HolKernel Parse boolLib; -infix THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL ## |->; -infixr -->; - (* --------------------------------------------------------------------- *) (* Create the theory. *) diff --git a/src/tfl/examples/sorting/partitionScript.sml b/src/tfl/examples/sorting/partitionScript.sml index 1c7bf534c5..b2389422cf 100644 --- a/src/tfl/examples/sorting/partitionScript.sml +++ b/src/tfl/examples/sorting/partitionScript.sml @@ -7,10 +7,6 @@ app load ["bossLib", "numLib"]; open HolKernel Parse boolLib numLib bossLib listTheory; -infix ## |-> THEN THENL THENC ORELSE ORELSEC THEN_TCL ORELSE_TCL; -infixr 3 -->; -infix 8 by; - val _ = new_theory "partition"; (*--------------------------------------------------------------------------- diff --git a/src/tfl/examples/sorting/permScript.sml b/src/tfl/examples/sorting/permScript.sml index 37af5ccd8e..5bc52e099b 100644 --- a/src/tfl/examples/sorting/permScript.sml +++ b/src/tfl/examples/sorting/permScript.sml @@ -3,9 +3,6 @@ ---------------------------------------------------------------------------*) open HolKernel Parse boolLib bossLib listTheory; -infix THEN THENL |-> ; -infix 8 by; - val _ = new_theory "perm"; (*---------------------------------------------------------------------------* diff --git a/src/tfl/examples/sorting/sortingScript.sml b/src/tfl/examples/sorting/sortingScript.sml index 5049f5b2aa..57006bc889 100644 --- a/src/tfl/examples/sorting/sortingScript.sml +++ b/src/tfl/examples/sorting/sortingScript.sml @@ -4,9 +4,6 @@ open HolKernel Parse boolLib bossLib listTheory permTheory; -infix THEN THENL -infix 8 by; - val MEM_APPEND_DISJ = Q.prove (`!x l1 l2. MEM x (APPEND l1 l2) = MEM x l1 \/ MEM x l2`, Induct_on `l1` THEN RW_TAC list_ss [APPEND,MEM] THEN PROVE_TAC[]); diff --git a/src/transfer/transferLib.sml b/src/transfer/transferLib.sml index add719d892..2756530047 100644 --- a/src/transfer/transferLib.sml +++ b/src/transfer/transferLib.sml @@ -6,7 +6,7 @@ open transferTheory FullUnify simpLib open pred_setTheory val op $ = Portable.$ -infix ++ +infix ++ |> ~> +++ type config = {cleftp:bool,force_imp:bool,hints:string list} @@ -31,7 +31,6 @@ fun relconstraint_tm t = | _ => false fun is_relconstraint t = relconstraint_tm (rator t) -infix ~> +++ fun sq ~> f = seq.flatten (seq.map f sq) fun sqreturn x = seq.fromList [x] fun sq1 +++ sq2 = diff --git a/tools/Holmake/HM_Core_Cline.sig b/tools/Holmake/HM_Core_Cline.sig index e840c80d90..88614b890a 100644 --- a/tools/Holmake/HM_Core_Cline.sig +++ b/tools/Holmake/HM_Core_Cline.sig @@ -10,6 +10,7 @@ type t = { includes : string list, interactive : bool, jobs : int, + json : bool, keep_going : bool, no_action : bool, no_hmakefile : bool, diff --git a/tools/Holmake/HM_Core_Cline.sml b/tools/Holmake/HM_Core_Cline.sml index 918433937c..b5072021e4 100644 --- a/tools/Holmake/HM_Core_Cline.sml +++ b/tools/Holmake/HM_Core_Cline.sml @@ -3,12 +3,12 @@ struct local open FunctionalRecordUpdate - fun makeUpdateT z = makeUpdate23 z + fun makeUpdateT z = makeUpdate24 z in fun updateT z = let fun from debug do_logging fast help hmakefile holdir includes - interactive jobs keep_going no_action no_hmakefile no_lastmaker_check - no_overlay + interactive jobs json keep_going no_action no_hmakefile + no_lastmaker_check no_overlay no_preexecs no_prereqs opentheory quiet quit_on_failure rebuild_deps recursive_build recursive_clean verbose = @@ -16,7 +16,7 @@ fun updateT z = let debug = debug, do_logging = do_logging, fast = fast, help = help, hmakefile = hmakefile, holdir = holdir, includes = includes, interactive = interactive, jobs = jobs, - keep_going = keep_going, + json = json, keep_going = keep_going, no_action = no_action, no_hmakefile = no_hmakefile, no_lastmaker_check = no_lastmaker_check, no_overlay = no_overlay, no_preexecs = no_preexecs, no_prereqs = no_prereqs, @@ -26,17 +26,16 @@ fun updateT z = let recursive_clean = recursive_clean, verbose = verbose } fun from' verbose recursive_clean recursive_build rebuild_deps quit_on_failure - quiet opentheory - no_prereqs no_preexecs + quiet opentheory no_prereqs no_preexecs no_overlay no_lastmaker_check no_hmakefile no_action keep_going - jobs interactive + json jobs interactive includes holdir hmakefile help fast do_logging debug = { debug = debug, do_logging = do_logging, fast = fast, help = help, hmakefile = hmakefile, holdir = holdir, includes = includes, interactive = interactive, jobs = jobs, - keep_going = keep_going, + json = json, keep_going = keep_going, no_action = no_action, no_hmakefile = no_hmakefile, no_lastmaker_check = no_lastmaker_check, no_overlay = no_overlay, no_preexecs = no_preexecs, no_prereqs = no_prereqs, @@ -46,14 +45,14 @@ fun updateT z = let recursive_clean = recursive_clean, verbose = verbose } fun to f {debug, do_logging, fast, help, hmakefile, holdir, - includes, interactive, jobs, keep_going, no_action, no_hmakefile, - no_lastmaker_check, + includes, interactive, jobs, json, keep_going, no_action, + no_hmakefile, no_lastmaker_check, no_overlay, no_preexecs, no_prereqs, opentheory, quiet, quit_on_failure, rebuild_deps, recursive_build, recursive_clean, verbose} = f debug do_logging fast help hmakefile holdir includes - interactive jobs keep_going no_action no_hmakefile no_lastmaker_check - no_overlay no_preexecs + interactive jobs json keep_going no_action no_hmakefile + no_lastmaker_check no_overlay no_preexecs no_prereqs opentheory quiet quit_on_failure rebuild_deps recursive_build recursive_clean verbose in @@ -76,6 +75,7 @@ type t = { includes : string list, interactive : bool, jobs : int, + json : bool, keep_going : bool, no_action : bool, no_hmakefile : bool, @@ -103,6 +103,7 @@ val default_core_options : t = includes = [], interactive = false, jobs = 4, + json = false, keep_going = false, no_action = false, no_hmakefile = false, @@ -203,6 +204,8 @@ val core_option_descriptions = [ desc = ReqArg (cons_includes, "directory") }, { help = "max number of parallel jobs", long = ["jobs"], short = "j", desc = ReqArg (change_jobs, "n") }, + { help = "emit JSON of dep. graph (then stop)", long = ["json"], + short = "", desc = mkBoolT #json }, { help = "try to build all targets", long = ["keep-going"], short = "k", desc = mkBoolT #keep_going }, { help = "enable time logging", long = ["logging"], short = "", diff --git a/tools/Holmake/HM_DepGraph.sig b/tools/Holmake/HM_DepGraph.sig index 161fdf148f..12aa12b37c 100644 --- a/tools/Holmake/HM_DepGraph.sig +++ b/tools/Holmake/HM_DepGraph.sig @@ -47,6 +47,7 @@ sig val find_runnable : 'a t -> (node * 'a nodeInfo) option val toString : 'a t -> string + val toJSONString : 'a t -> string val postmortem : Holmake_tools.output_functions -> OS.Process.status * 'a t -> OS.Process.status diff --git a/tools/Holmake/HM_DepGraph.sml b/tools/Holmake/HM_DepGraph.sml index 356eeb1f7a..bc25834d34 100644 --- a/tools/Holmake/HM_DepGraph.sml +++ b/tools/Holmake/HM_DepGraph.sml @@ -35,6 +35,11 @@ datatype command = NoCmd | SomeCmd of string | BuiltInCmd of builtincmd * Holmake_tools.include_info + +fun command_toString NoCmd = "" + | command_toString (SomeCmd s) = s + | command_toString (BuiltInCmd(bic,_)) = bic_toString bic + type dir = Holmake_tools.hmdir.t type 'a nodeInfo = { target : dep, status : target_status, extra : 'a, command : command, phony : bool, @@ -171,6 +176,42 @@ fun nodeInfo_toString (nI : 'a nodeInfo) = | NoCmd => "") end +datatype 'a applist = apNIL | apSING of 'a | ++ of ('a applist * 'a applist) +infix ++ +fun ap2list apNIL A = A + | ap2list (apSING x) A = x::A + | ap2list (a ++ b) A = ap2list a (ap2list b A) + + +fun pr_list s [] = apNIL + | pr_list s [x] = x + | pr_list s (x::xs) = x ++ apSING s ++ pr_list s xs + +fun nodeInfo_toJSON (n, nI : 'a nodeInfo) = + let + open Holmake_tools + val {target,status,command,dependencies,seqnum,phony,dir,...} = nI + fun field fnm f v = apSING (" \"" ^ fnm ^ "\" : " ^ f v) + fun quote f x = "\"" ^ f x ^ "\"" + in + apSING "{\n" ++ + pr_list ",\n" [ + field "node_id" Int.toString n, + field "target" (quote hm_target.toString) target, + field "seqnum" Int.toString seqnum, + field "phony" Bool.toString phony, + field "dependencies" + (fn ds => + "[" ^ + String.concatWith ", " (map (Int.toString o #1) ds) ^ "]") + dependencies, + field "command" (fn c => "\"" ^ command_toString c ^ "\"") command, + field "dir" (quote hmdir.toString) dir, + field "needs_rebuild" (fn s => Bool.toString (s <> Succeeded)) status + ] ++ + apSING "\n}" + end + fun mkneeded tgts g = let fun setneeded f n g = updnode(n,f{needed=true}) g @@ -251,6 +292,17 @@ fun toString g = String.concatWith ",\n " (map prNode others) ^ "\n}" end +fun toJSONString g = + let + val ns = listNodes g + val ap = apSING "[\n" ++ + pr_list ",\n" (map nodeInfo_toJSON ns) ++ + apSING "\n]" + val ss = ap2list ap [] + in + String.concat ss + end + fun postmortem (outs : Holmake_tools.output_functions) (status,g) = let val pr = tgt_toString diff --git a/tools/Holmake/HolLex b/tools/Holmake/HolLex index 14a460cbfe..f05a6b86dc 100644 --- a/tools/Holmake/HolLex +++ b/tools/Holmake/HolLex @@ -214,7 +214,7 @@ fullquotebegin = "``" | "\226\128\156"; fullquoteend = "``" | "\226\128\157"; quotebegin = "`" | "\226\128\152"; quoteend = "`" | "\226\128\153"; -Theorempfx = ("Theorem"|"Triviality"){ws}+{alphaMLid}({ws}*"["{alphaMLid_list}"]")?; +Theorempfx = ("Theorem"|"Triviality"){ws}+{alphaMLid}({ws}*"["{optallws}{defn_attribute_list}"]")?; Quote_eqnpfx = "Quote"{ws}+{alphaMLid}{ws}*"="{ws}*{QUALalphaMLid}{ws}*":"; Quote_pfx = "Quote"{ws}+{QUALalphaMLid}{ws}*":"; HOLconjunction = "/\092" | "\226\136\167"; diff --git a/tools/Holmake/Holmake.sml b/tools/Holmake/Holmake.sml index 73d7c359da..115451de29 100644 --- a/tools/Holmake/Holmake.sml +++ b/tools/Holmake/Holmake.sml @@ -196,6 +196,13 @@ fun extend_with_cline_vars env = end +(* ---------------------------------------------------------------------- + get_hmf : unit -> "holmakefile data" (as per ReadHMF) + + Utility function to get the Holmakefile in the current directory, but + using a cache so that any given file is only ever read once. + ---------------------------------------------------------------------- *) + local open hm_target val base = extend_with_cline_vars (read_holpathdb()) @@ -301,6 +308,7 @@ val (outputfns as {warn,tgtfatal,diag,info,chatty,info_inline,info_inline_end})= val do_logging_flag = #do_logging coption_value val no_lastmakercheck = #no_lastmaker_check coption_value val show_usage = #help coption_value +val show_json = #json coption_value val quit_on_failure = #quit_on_failure coption_value val toplevel_no_prereqs = #no_prereqs coption_value val toplevel_no_overlay = #no_overlay coption_value @@ -309,6 +317,10 @@ val cline_always_rebuild_deps = #rebuild_deps coption_value val cline_nobuild = #no_action coption_value val cline_recursive_build = #recursive_build coption_value val cline_recursive_clean = #recursive_clean coption_value +val scan_output_functions = + if show_json andalso chattiness_level coption_value <= 1 then + quieten_info outputfns + else outputfns (* make the cline includes = [] so that these are only looked at once (when the cline_additional_includes value is folded into dirinfo values @@ -429,7 +441,8 @@ fun 'a recursively getnewincs dsopt {outputfns,verb,hm,dirinfo,dir,data} = let val {incdirmap,visited,ancestors} = dirinfo : dirinfo val hm : 'a hmfold = hm - val {warn,diag,info,chatty,...} : output_functions = outputfns + val {warn,diag,info,chatty,info_inline_end, info_inline,...} : output_functions = + outputfns val {includes=incset, preincludes = preincset} = getnewincs dir val incdirmap = incdirmap |> extend_idmap dir {incs = incset, pres = preincset} @@ -1086,7 +1099,7 @@ fun create_complete_graph cline_incs idm = val d = hmdir.curdir() val {data = g, incdirmap, visited, ...} = recursively getnewincs (SOME cline_incs) { - outputfns = outputfns, verb = "Scanning", + outputfns = scan_output_functions, verb = "Scanning", hm=extend_graph_in_dir, dirinfo={incdirmap=idm, visited = Binaryset.empty hmdir.compare, ancestors = [original_dir]}, @@ -1095,9 +1108,10 @@ fun create_complete_graph cline_incs idm = } val numScanned = Binaryset.numItems visited val _ = if numScanned > 1 then - (#info_inline outputfns ("Scanned " ^ Int.toString numScanned ^ - " directories"); - #info_inline_end outputfns()) + (#info_inline scan_output_functions + ("Scanned " ^ Int.toString numScanned ^ + " directories"); + #info_inline_end scan_output_functions()) else () val diag = diag "builddepgraph" in @@ -1205,7 +1219,10 @@ fun work() = in OS.Process.success end - else + else if show_json then ( + info (HM_DepGraph.toJSONString depgraph); + OS.Process.exit OS.Process.success + ) else (* actually build default targets *) postmortem outputfns (build_graph depgraph) handle e => die ("Exception: "^General.exnMessage e) end diff --git a/tools/Holmake/Holmake_tools.sig b/tools/Holmake/Holmake_tools.sig index f5995de98d..d5e63bb5b8 100644 --- a/tools/Holmake/Holmake_tools.sig +++ b/tools/Holmake/Holmake_tools.sig @@ -80,6 +80,11 @@ sig chatty : string -> unit, tgtfatal : string -> unit, diag : string -> (unit -> string) -> unit} + type 'a ofn_update = ('a -> 'a) -> output_functions -> output_functions + val fupd_info_inline : (string -> unit) ofn_update + val fupd_info_inline_end : (unit -> unit) ofn_update + val fupd_info : (string -> unit) ofn_update + val quieten_info : output_functions -> output_functions (* 0 : quiet, 1 : normal, 2 : chatty, 3 : everything + debug info *) val output_functions : {chattiness:int, diff --git a/tools/Holmake/Holmake_tools.sml b/tools/Holmake/Holmake_tools.sml index f67bc1101e..27dc404c7f 100644 --- a/tools/Holmake/Holmake_tools.sml +++ b/tools/Holmake/Holmake_tools.sml @@ -216,6 +216,19 @@ type output_functions = {warn : string -> unit, chatty : string -> unit, tgtfatal : string -> unit, diag : string -> (unit -> string) -> unit} +type 'a ofn_update = ('a -> 'a) -> output_functions -> output_functions +fun fupd_info_inline f {warn,info,info_inline,info_inline_end,chatty,tgtfatal,diag} = + {warn = warn,info = info, info_inline_end = info_inline_end, info_inline = f info_inline, + chatty = chatty, tgtfatal = tgtfatal, diag = diag} +fun fupd_info_inline_end f {warn,info,info_inline,info_inline_end,chatty,tgtfatal,diag} = + {warn = warn,info = info, info_inline_end = f info_inline_end, info_inline = info_inline, + chatty = chatty, tgtfatal = tgtfatal, diag = diag} +fun fupd_info f {warn,info,info_inline,info_inline_end,chatty,tgtfatal,diag} = + {warn = warn,info = f info, info_inline_end = info_inline_end, info_inline = info_inline, + chatty = chatty, tgtfatal = tgtfatal, diag = diag} +fun quieten_info ofns = ofns |> fupd_info_inline_end (K (K ())) + |> fupd_info_inline (K (K ())) + |> fupd_info (K (K ())) fun die_with message = let in diff --git a/tools/Holmake/tests/quote-filter/expected-mosml b/tools/Holmake/tests/quote-filter/expected-mosml index edbf724b80..da1689e5c9 100644 --- a/tools/Holmake/tests/quote-filter/expected-mosml +++ b/tools/Holmake/tests/quote-filter/expected-mosml @@ -93,3 +93,12 @@ val x = 92 + 10; val y = 11 val s = "" val s' = "foo.sml" (* #(FILE) *) + +val thm = Q.store_thm_at (DB_dtype.mkloc ("foo.sml", 15, true)) ("thm[foo, simp , key = val1 val2]", [QUOTE " (*#loc 15 42*)\n\ +\ p /\\ q"], (fn HOL__GOAL__foo => ( + + cheat +) HOL__GOAL__foo)); + +val thm2 = boolLib.save_thm("thm2[foo, simp , key = val1 val2]", exp +); \ No newline at end of file diff --git a/tools/Holmake/tests/quote-filter/expected-poly b/tools/Holmake/tests/quote-filter/expected-poly index 9c2fcad688..230eaf478e 100644 --- a/tools/Holmake/tests/quote-filter/expected-poly +++ b/tools/Holmake/tests/quote-filter/expected-poly @@ -93,3 +93,12 @@ val x = 92 + 10; val y = 11 val s = "" val s' = "foo.sml" (* #(FILE) *) + +val thm = Q.store_thm_at (DB_dtype.mkloc ("foo.sml", 15, true)) ("thm[foo, simp , key = val1 val2]", [QUOTE " (*#loc 15 42*)\n\ +\ p /\\ q"], (fn HOL__GOAL__foo => ( + + cheat +) HOL__GOAL__foo)); + +val thm2 = boolLib.save_thm("thm2[foo, simp , key = val1 val2]", exp +); \ No newline at end of file diff --git a/tools/Holmake/tests/quote-filter/input b/tools/Holmake/tests/quote-filter/input index 28f1c65729..d9835238d6 100644 --- a/tools/Holmake/tests/quote-filter/input +++ b/tools/Holmake/tests/quote-filter/input @@ -93,3 +93,11 @@ val x = #(LINE) + 10; val #(LINE=11) y = #(LINE) val s = #(FILE) #(FILE=foo.sml) val s' = #(FILE) (* #(FILE) *) + +Theorem thm[foo, simp , key = val1 val2]: + p /\ q +Proof + cheat +QED + +Theorem thm2[foo, simp , key = val1 val2] = exp