author | wenzelm |
Sat, 04 Jan 2025 23:20:05 +0100 | |
changeset 81724 | ed5a05b06ee0 |
parent 81717 | 114449035ec6 |
child 81731 | 3af5379aac0e |
permissions | -rw-r--r-- |
71099 | 1 |
(*:maxLineLen=78:*) |
2 |
||
3 |
theory Phabricator |
|
4 |
imports Base |
|
5 |
begin |
|
6 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
7 |
chapter \<open>Phabricator / Phorge server setup \label{ch:phabricator}\<close> |
71099 | 8 |
|
9 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
10 |
Phorge\<^footnote>\<open>\<^url>\<open>https://phorge.it\<close>\<close> is an open-source product to support the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
11 |
development process of complex software projects (open or closed ones). It |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
12 |
is a community fork to replace the former |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
13 |
Phabricator\<^footnote>\<open>\<^url>\<open>https://www.phacility.com/phabricator\<close>\<close> project, which is |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
14 |
now inactive. Subsequently, the product name is always \<^emph>\<open>Phorge\<close> instead of |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
15 |
\<^emph>\<open>Phabricator\<close>, but files and other formal names usually refer to |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
16 |
\<^verbatim>\<open>phabricator\<close>. |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
17 |
|
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
18 |
Following the original tradition of Phabricator, almost everything in Phorge |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
19 |
is a bit different and unusual. The official project description is: |
71099 | 20 |
|
21 |
\begin{quote} |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
22 |
Your opinionated Free/Libre and Open Source, community driven platform |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
23 |
for collaborating, managing, organizing and reviewing software projects. |
71099 | 24 |
\end{quote} |
25 |
||
26 |
Ongoing changes and discussions about changes are maintained uniformly |
|
27 |
within a MySQL database. There are standard connections to major version |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
28 |
control systems: \<^bold>\<open>Subversion\<close>, \<^bold>\<open>Mercurial\<close>, \<^bold>\<open>Git\<close>. So Phorge offers a |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
29 |
counter-model to trends of monopoly and centralized version control, |
71099 | 30 |
especially due to Microsoft's Github and Atlassian's Bitbucket. |
31 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
32 |
A notable public instance of Phorge is running on \<^url>\<open>https://gitpull.it\<close>. |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
33 |
Independent \<^emph>\<open>self-hosting\<close> requires an old-school LAMP server (Linux, |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
34 |
Apache, MySQL, PHP): a cheap virtual machine on the Net is sufficient, there |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
35 |
is no need for special ``cloud'' providers. So it is feasible to remain the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
36 |
master of your virtual home, according to the principle ``to own all your |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
37 |
data''. Thus Phorge is similar to the well-known |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
38 |
Nextcloud\<^footnote>\<open>\<^url>\<open>https://nextcloud.com\<close>\<close> server product, concerning both the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
39 |
technology and sociology. |
71099 | 40 |
|
41 |
\<^medskip> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
42 |
Initial Phorge configuration requires many details to be done right. |
71296 | 43 |
Isabelle provides some command-line tools to help with the setup, and |
44 |
afterwards Isabelle support is optional: it is possible to run and maintain |
|
45 |
the server, without requiring the somewhat bulky Isabelle distribution |
|
46 |
again. |
|
71322 | 47 |
|
48 |
\<^medskip> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
49 |
Assuming an existing installation of Phorge, the Isabelle command-line tool |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
50 |
@{tool hg_setup} (\secref{sec:hg-setup}) helps to create new repositories or |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
51 |
to migrate old ones. In particular, this avoids the lengthy sequence of |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
52 |
clicks in Phorge to make a new private repository with hosting on the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
53 |
server. (Phorge is a software project management platform, where initial |
71322 | 54 |
repository setup happens rarely in practice.) |
71099 | 55 |
\<close> |
56 |
||
57 |
||
58 |
section \<open>Quick start\<close> |
|
59 |
||
60 |
text \<open> |
|
81724 | 61 |
The starting point is a fresh installation of \<^bold>\<open>Ubuntu 22.04 or 24.04 |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
62 |
LTS\<close>\<^footnote>\<open>\<^url>\<open>https://ubuntu.com/download\<close>\<close>: these versions are mandatory due to |
71131 | 63 |
subtle dependencies on system packages and configuration that is assumed by |
64 |
the Isabelle setup tool. |
|
71099 | 65 |
|
66 |
For production use, a proper \<^emph>\<open>Virtual Server\<close> or \<^emph>\<open>Root Server\<close> product |
|
67 |
from a hosting provider will be required, including an Internet Domain Name |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
68 |
(\secref{sec:phorge-domain}). Initial experimentation also works on a local |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
69 |
host, e.g.\ via VirtualBox\<^footnote>\<open>\<^url>\<open>https://www.virtualbox.org\<close>\<close>. The proforma |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
70 |
domain \<^verbatim>\<open>localhost\<close> is used by default: it maps arbitrary subdomains to the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
71 |
usual \<^verbatim>\<open>localhost\<close> address. This allows to use e.g. |
79482 | 72 |
\<^verbatim>\<open>http://phabricator-vcs.localhost\<close> for initial setup as described below. |
71099 | 73 |
|
74 |
All administrative commands need to be run as \<^verbatim>\<open>root\<close> user (e.g.\ via |
|
71131 | 75 |
\<^verbatim>\<open>sudo\<close>). Note that Isabelle refers to user-specific configuration in the |
76 |
user home directory via @{setting ISABELLE_HOME_USER} |
|
77 |
(\secref{sec:settings}); that may be different or absent for the root user |
|
78 |
and thus cause confusion. |
|
71099 | 79 |
\<close> |
80 |
||
81 |
||
82 |
subsection \<open>Initial setup\<close> |
|
83 |
||
84 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
85 |
Isabelle can manage multiple named Phorge installations: this allows to |
71131 | 86 |
separate administrative responsibilities, e.g.\ different approaches to user |
87 |
management for different projects. Subsequently we always use the default |
|
71290 | 88 |
name ``\<^verbatim>\<open>vcs\<close>'': the name will appear in file and directory locations, |
89 |
internal database names and URLs. |
|
71099 | 90 |
|
71131 | 91 |
The initial setup works as follows (with full Linux package upgrade): |
71099 | 92 |
|
71280
5a2033fc8f3d
avoid odd (harmless) problem with Mercurial 4.5.3 provided by Ubuntu 18.04 on first push: "couldn't write revision branch cache names";
wenzelm
parents:
71279
diff
changeset
|
93 |
@{verbatim [display] \<open> isabelle phabricator_setup -U -M:\<close>} |
71099 | 94 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
95 |
After installing many packages, cloning the Phorge distribution, |
71131 | 96 |
initializing the MySQL database and Apache, the tool prints an URL for |
71290 | 97 |
further configuration. Now the following needs to be provided by the web |
98 |
interface. |
|
71099 | 99 |
|
71131 | 100 |
\<^item> An initial user that will get administrator rights. There is no need to |
101 |
create a special \<^verbatim>\<open>admin\<close> account. Instead, a regular user that will take |
|
102 |
over this responsibility can be used here. Subsequently we assume that |
|
103 |
user \<^verbatim>\<open>makarius\<close> becomes the initial administrator. |
|
71099 | 104 |
|
105 |
\<^item> An \<^emph>\<open>Auth Provider\<close> to manage user names and passwords. None is provided |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
106 |
by default, and Phorge points out this omission prominently in its |
71099 | 107 |
overview of \<^emph>\<open>Setup Issues\<close>: following these hints quickly leads to the |
108 |
place where a regular \<^emph>\<open>Username/Password\<close> provider can be added. |
|
109 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
110 |
Alternatively, Phorge can delegate the responsibility of |
71131 | 111 |
authentication to big corporations like Google and Facebook, but these can |
71290 | 112 |
be easily ignored. Genuine self-hosting means to manage users directly, |
113 |
without outsourcing of authentication. |
|
71099 | 114 |
|
71131 | 115 |
\<^item> A proper password for the administrator can now be set, e.g.\ by the |
116 |
following command: |
|
71099 | 117 |
|
71131 | 118 |
@{verbatim [display] \<open> isabelle phabricator bin/auth recover makarius\<close>} |
71099 | 119 |
|
71131 | 120 |
The printed URL gives access to a login and password dialog in the web |
121 |
interface. |
|
71099 | 122 |
|
71131 | 123 |
Any further users will be able to provide a password directly, because the |
124 |
Auth Provider is already active. |
|
125 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
126 |
\<^item> The list of Phorge \<^bold>\<open>Setup Issues\<close> should be studied with some care, to |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
127 |
make sure that no serious problems are remaining. For example, the request |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
128 |
to lock the configuration can be fulfilled as follows: |
71099 | 129 |
|
71131 | 130 |
@{verbatim [display] \<open> isabelle phabricator bin/auth lock\<close>} |
71099 | 131 |
|
71131 | 132 |
\<^medskip> A few other Setup Issues might be relevant as well, e.g.\ the timezone |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
133 |
of the server. Some more exotic points can be ignored: Phorge provides |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
134 |
careful explanations about what it thinks could be wrong, while leaving |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
135 |
some room for interpretation. |
71099 | 136 |
\<close> |
137 |
||
138 |
||
139 |
subsection \<open>Mailer configuration\<close> |
|
140 |
||
141 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
142 |
The next important thing is messaging: Phorge needs to be able to |
71131 | 143 |
communicate with users on its own account, e.g.\ to reset passwords. The |
144 |
documentation has many variations on \<^emph>\<open>Configuring Outbound |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
145 |
Email\<close>\<^footnote>\<open>\<^url>\<open>https://we.phorge.it/book/phorge/article/configuring_outbound_email\<close>\<close>, |
71131 | 146 |
but a conventional SMTP server with a dedicated \<^verbatim>\<open>phabricator\<close> user is |
147 |
sufficient. There is no need to run a separate mail server on the |
|
148 |
self-hosted Linux machine: hosting providers often include such a service |
|
149 |
for free, e.g.\ as part of a web-hosting package. As a last resort it is |
|
150 |
also possible to use a corporate service like Gmail, but such dependency |
|
151 |
dilutes the whole effort of self-hosting. |
|
71099 | 152 |
|
153 |
\<^medskip> |
|
154 |
Mailer configuration requires a few command-line invocations as follows: |
|
155 |
||
156 |
@{verbatim [display] \<open> isabelle phabricator_setup_mail\<close>} |
|
157 |
||
75161
95612f330c93
misc tuning based on comments by Heiko Eißfeldt;
wenzelm
parents:
73534
diff
changeset
|
158 |
\<^noindent> This generates a JSON template file for the mail account details. |
71131 | 159 |
After editing that, the subsequent command will add and test it with |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
160 |
Phorge: |
71099 | 161 |
|
162 |
@{verbatim [display] \<open> isabelle phabricator_setup_mail -T makarius\<close>} |
|
163 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
164 |
This tells Phorge to send a message to the administrator created |
71131 | 165 |
before; the output informs about success or errors. |
166 |
||
167 |
The mail configuration process can be refined and repeated until it works |
|
168 |
properly: host name, port number, protocol etc.\ all need to be correct. The |
|
169 |
\<^verbatim>\<open>key\<close> field in the JSON file identifies the name of the configuration that |
|
71099 | 170 |
will be overwritten each time, when taking over the parameters via |
171 |
\<^verbatim>\<open>isabelle phabricator_setup_mail\<close>. |
|
172 |
||
173 |
\<^medskip> |
|
71290 | 174 |
The effective mail configuration can be queried like this: |
71099 | 175 |
|
71103 | 176 |
@{verbatim [display] \<open> isabelle phabricator bin/config get cluster.mailers\<close>} |
177 |
\<close> |
|
178 |
||
179 |
||
71132 | 180 |
subsection \<open>SSH configuration\<close> |
181 |
||
182 |
text \<open> |
|
71323 | 183 |
SSH configuration is important to access hosted repositories with public-key |
184 |
authentication. It is done by a separate tool, because it affects the |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
185 |
operating-system and all installations of Phorge simultaneously. |
71132 | 186 |
|
71290 | 187 |
The subsequent configuration is convenient (and ambitious): it takes away |
188 |
the standard port 22 from the operating system and assigns it to |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
189 |
Isabelle/Phorge. |
71132 | 190 |
|
191 |
@{verbatim [display] \<open> isabelle phabricator_setup_ssh -p 22 -q 222\<close>} |
|
192 |
||
193 |
Afterwards, remote login to the server host needs to use that alternative |
|
71290 | 194 |
port 222. If there is a problem connecting again, the administrator can |
195 |
usually access a remote console via some web interface of the virtual server |
|
71267 | 196 |
provider. |
71132 | 197 |
|
198 |
\<^medskip> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
199 |
The following alternative is more modest: it uses port 2222 for Phorge, and |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
200 |
retains port 22 for the operating system. |
71132 | 201 |
|
202 |
@{verbatim [display] \<open> isabelle phabricator_setup_ssh -p 2222 -q 22\<close>} |
|
203 |
||
204 |
\<^medskip> |
|
205 |
The tool can be invoked multiple times with different parameters; ports are |
|
206 |
changed back and forth each time and services restarted. |
|
207 |
\<close> |
|
208 |
||
209 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
210 |
subsection \<open>Internet domain name and HTTPS configuration \label{sec:phorge-domain}\<close> |
71278 | 211 |
|
212 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
213 |
So far the Phorge server has been accessible only on \<^verbatim>\<open>localhost\<close>. Proper |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
214 |
configuration of a public Internet domain name (with HTTPS certificate from |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
215 |
\<^emph>\<open>Let's Encrypt\<close>) works as follows. |
71278 | 216 |
|
217 |
\<^item> Register a subdomain (e.g.\ \<^verbatim>\<open>vcs.example.org\<close>) as an alias for the IP |
|
218 |
address of the underlying Linux host. This usually works by some web |
|
219 |
interface of the hosting provider to edit DNS entries; it might require |
|
220 |
some time for updated DNS records to become publicly available. |
|
221 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
222 |
\<^item> Edit the Phorge website configuration file in |
71278 | 223 |
\<^path>\<open>/etc/apache2/sites-available/\<close> to specify \<^verbatim>\<open>ServerName\<close> and |
224 |
\<^verbatim>\<open>ServerAdmin\<close> like this: @{verbatim [display] \<open> ServerName vcs.example.org |
|
225 |
ServerAdmin webmaster@example.org\<close>} |
|
226 |
||
227 |
Then reload (or restart) Apache like this: |
|
228 |
@{verbatim [display] \<open> systemctl reload apache2\<close>} |
|
229 |
||
230 |
\<^item> Install \<^verbatim>\<open>certbot\<close> from \<^url>\<open>https://certbot.eff.org\<close> following the |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
231 |
description for Apache and Ubuntu Linux. Run \<^verbatim>\<open>certbot\<close> interactively and |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
232 |
let it operate on the domain \<^verbatim>\<open>vcs.example.org\<close>. |
71278 | 233 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
234 |
\<^item> Inform Phorge about its new domain name like this: |
71290 | 235 |
@{verbatim [display] \<open> isabelle phabricator bin/config set \ |
71278 | 236 |
phabricator.base-uri https://vcs.example.org\<close>} |
237 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
238 |
\<^item> Visit the website \<^verbatim>\<open>https://vcs.example.org\<close> and configure Phorge |
71290 | 239 |
as described before. The following options are particularly relevant for a |
240 |
public website: |
|
71278 | 241 |
|
242 |
\<^item> \<^emph>\<open>Auth Provider / Username/Password\<close>: disable \<^emph>\<open>Allow Registration\<close> to |
|
71290 | 243 |
avoid uncontrolled registrants; users can still be invited via email |
244 |
instead. |
|
71278 | 245 |
|
246 |
\<^item> Enable \<^verbatim>\<open>policy.allow-public\<close> to allow read-only access to resources, |
|
247 |
without requiring user registration. |
|
71390
5e7ba6aa85d7
more documentation: odd option for special situations;
wenzelm
parents:
71365
diff
changeset
|
248 |
|
5e7ba6aa85d7
more documentation: odd option for special situations;
wenzelm
parents:
71365
diff
changeset
|
249 |
\<^item> Adjust \<^verbatim>\<open>phabricator.cookie-prefix\<close> for multiple installations with |
5e7ba6aa85d7
more documentation: odd option for special situations;
wenzelm
parents:
71365
diff
changeset
|
250 |
overlapping domains (see also the documentation of this configuration |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
251 |
option within Phorge). |
71278 | 252 |
\<close> |
253 |
||
254 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
255 |
section \<open>Global data storage and backups \label{sec:phorge-backup}\<close> |
71271 | 256 |
|
257 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
258 |
The global state of a Phorge installation consists of two main parts: |
71271 | 259 |
|
260 |
\<^enum> The \<^emph>\<open>root directory\<close> according to |
|
261 |
\<^path>\<open>/etc/isabelle-phabricator.conf\<close> or \<^verbatim>\<open>isabelle phabricator -l\<close>: it |
|
262 |
contains the main PHP program suite with administrative tools, and some |
|
263 |
configuration files. The default setup also puts hosted repositories here |
|
264 |
(subdirectory \<^verbatim>\<open>repo\<close>). |
|
265 |
||
266 |
\<^enum> Multiple \<^emph>\<open>MySQL databases\<close> with a common prefix derived from the |
|
267 |
installation name --- the same name is used as database user name. |
|
268 |
||
71329 | 269 |
The root user may invoke \<^verbatim>\<open>/usr/local/bin/isabelle-phabricator-dump\<close> to |
270 |
create a complete database dump within the root directory. Afterwards it is |
|
271 |
sufficient to make a conventional \<^bold>\<open>file-system backup\<close> of everything. To |
|
71271 | 272 |
restore the database state, see the explanations on \<^verbatim>\<open>mysqldump\<close> in |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
273 |
\<^url>\<open>https://we.phorge.it/book/phorge/article/configuring_backups\<close>; some |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
274 |
background information is in |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
275 |
\<^url>\<open>https://we.phorge.it/book/flavor/article/so_many_databases\<close>. |
71279 | 276 |
|
277 |
\<^medskip> The following command-line tools are particularly interesting for advanced |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
278 |
database maintenance (within the Phorge root directory that is traditionally |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
279 |
called \<^verbatim>\<open>phabricator\<close>): |
71279 | 280 |
@{verbatim [display] \<open> phabricator/bin/storage help dump |
281 |
phabricator/bin/storage help shell |
|
282 |
phabricator/bin/storage help destroy |
|
283 |
phabricator/bin/storage help renamespace\<close>} |
|
284 |
||
285 |
For example, copying a database snapshot from one installation to another |
|
286 |
works as follows. Run on the first installation root directory: |
|
287 |
||
288 |
@{verbatim [display] \<open> phabricator/bin/storage dump > dump1.sql |
|
71290 | 289 |
phabricator/bin/storage renamespace --from phabricator_vcs \ |
71279 | 290 |
--to phabricator_xyz --input dump1.sql --output dump2.sql\<close>} |
291 |
||
79485 | 292 |
Then run on the second installation root directory: |
71279 | 293 |
@{verbatim [display] \<open> phabricator/bin/storage destroy |
71290 | 294 |
phabricator/bin/storage shell < .../dump2.sql\<close>} |
71279 | 295 |
|
296 |
Local configuration in \<^verbatim>\<open>phabricator/config/local/\<close> and hosted repositories |
|
297 |
need to be treated separately within the file-system. For the latter |
|
298 |
see also these tools: |
|
299 |
@{verbatim [display] \<open> phabricator/bin/repository help list-paths |
|
300 |
phabricator/bin/repository help move-paths\<close>} |
|
71271 | 301 |
\<close> |
302 |
||
303 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
304 |
section \<open>Upgrading Phorge installations\<close> |
71286 | 305 |
|
306 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
307 |
The Phorge community publishes a new stable version several times per year: |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
308 |
see also \<^url>\<open>https://we.phorge.it/w/changelog\<close>. There is no need to follow |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
309 |
updates on the spot, but it is a good idea to upgrade occasionally --- with |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
310 |
the usual care to avoid breaking a production system (see also |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
311 |
\secref{sec:phorge-backup} for database dump and backup). |
71286 | 312 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
313 |
The Isabelle/Phorge setup provides a convenience tool to upgrade all |
71286 | 314 |
installations uniformly: |
315 |
@{verbatim [display] \<open> /usr/local/bin/isabelle-phabricator-upgrade\<close>} |
|
316 |
||
317 |
This refers to the \<^verbatim>\<open>stable\<close> branch of the distribution repositories by |
|
318 |
default. Alternatively, it also possible to use the \<^verbatim>\<open>master\<close> like this: |
|
319 |
@{verbatim [display] \<open> /usr/local/bin/isabelle-phabricator-upgrade master\<close>} |
|
71296 | 320 |
|
321 |
\<^medskip> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
322 |
See \<^url>\<open>https://we.phorge.it/book/phorge/article/upgrading\<close> for further |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
323 |
explanations on Phorge upgrade. |
71286 | 324 |
\<close> |
325 |
||
326 |
||
71103 | 327 |
section \<open>Reference of command-line tools\<close> |
328 |
||
329 |
text \<open> |
|
330 |
The subsequent command-line tools usually require root user privileges on |
|
331 |
the underlying Linux system (e.g.\ via \<^verbatim>\<open>sudo bash\<close> to open a subshell, or |
|
71131 | 332 |
directly via \<^verbatim>\<open>sudo isabelle phabricator ...\<close>). |
71103 | 333 |
\<close> |
334 |
||
335 |
||
336 |
subsection \<open>\<^verbatim>\<open>isabelle phabricator\<close>\<close> |
|
337 |
||
338 |
text \<open> |
|
339 |
The @{tool_def phabricator} tool invokes a GNU bash command-line within the |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
340 |
Phorge home directory: |
71103 | 341 |
@{verbatim [display] |
342 |
\<open>Usage: isabelle phabricator [OPTIONS] COMMAND [ARGS...] |
|
343 |
||
344 |
Options are: |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
345 |
-l list available Phorge installations |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
346 |
-n NAME Phorge installation name (default: "vcs") |
71103 | 347 |
|
348 |
Invoke a command-line tool within the home directory of the named |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
349 |
Phorge installation.\<close>} |
71103 | 350 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
351 |
Isabelle/Phorge installations are registered in the global configuration |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
352 |
file \<^path>\<open>/etc/isabelle-phabricator.conf\<close>, with name and root directory |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
353 |
separated by colon (no extra whitespace). The home directory is the |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
354 |
subdirectory \<^verbatim>\<open>phabricator\<close> within the root. |
71103 | 355 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
356 |
\<^medskip> Option \<^verbatim>\<open>-l\<close> lists the available Phorge installations with name and root |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
357 |
directory --- without invoking a command. |
71103 | 358 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
359 |
Option \<^verbatim>\<open>-n\<close> selects the explicitly named Phorge installation. |
71103 | 360 |
\<close> |
361 |
||
362 |
||
363 |
subsubsection \<open>Examples\<close> |
|
364 |
||
365 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
366 |
Print the home directory of the Phorge installation: |
71132 | 367 |
@{verbatim [display] \<open> isabelle phabricator pwd\<close>} |
71103 | 368 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
369 |
Print some Phorge configuration information: |
71132 | 370 |
@{verbatim [display] \<open> isabelle phabricator bin/config get phabricator.base-uri\<close>} |
71103 | 371 |
|
372 |
The latter conforms to typical command templates seen in the original |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
373 |
Phorge documentation: |
71132 | 374 |
@{verbatim [display] \<open> phabricator/ $ ./bin/config get phabricator.base-uri\<close>} |
71103 | 375 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
376 |
Here the user is meant to navigate to the Phorge home manually, in |
71131 | 377 |
contrast to \<^verbatim>\<open>isabelle phabricator\<close> doing it automatically thanks to the |
378 |
global configuration \<^path>\<open>/etc/isabelle-phabricator.conf\<close>. |
|
71103 | 379 |
\<close> |
380 |
||
381 |
||
382 |
subsection \<open>\<^verbatim>\<open>isabelle phabricator_setup\<close>\<close> |
|
383 |
||
384 |
text \<open> |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
385 |
The @{tool_def phabricator_setup} tool installs a fresh Phorge instance |
81724 | 386 |
on Ubuntu 22.04 or 24.04 LTS: |
71103 | 387 |
@{verbatim [display] \<open>Usage: isabelle phabricator_setup [OPTIONS] |
388 |
||
389 |
Options are: |
|
71280
5a2033fc8f3d
avoid odd (harmless) problem with Mercurial 4.5.3 provided by Ubuntu 18.04 on first push: "couldn't write revision branch cache names";
wenzelm
parents:
71279
diff
changeset
|
390 |
-M SOURCE install Mercurial from source: local PATH, or URL, or ":" |
71103 | 391 |
-R DIR repository directory (default: "/var/www/phabricator-NAME/repo") |
392 |
-U full update of system packages before installation |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
393 |
-n NAME Phorge installation name (default: "vcs") |
71422
5d5be87330b5
allow to override repository versions at runtime;
wenzelm
parents:
71390
diff
changeset
|
394 |
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) |
71103 | 395 |
-r DIR installation root directory (default: "/var/www/phabricator-NAME") |
396 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
397 |
Install Phorge as LAMP application (Linux, Apache, MySQL, PHP). |
71103 | 398 |
|
399 |
The installation name (default: "vcs") is mapped to a regular |
|
400 |
Unix user; this is relevant for public SSH access.\<close>} |
|
401 |
||
71131 | 402 |
Installation requires Linux root permissions. All required packages are |
71103 | 403 |
installed automatically beforehand, this includes the Apache web server and |
404 |
the MySQL database engine. |
|
405 |
||
71131 | 406 |
Global configuration in \<^verbatim>\<open>/etc\<close> or a few other directories like \<^verbatim>\<open>/var/www\<close> |
407 |
uses name prefixes like \<^verbatim>\<open>isabelle-phabricator\<close> or \<^verbatim>\<open>phabricator\<close>. Local |
|
71103 | 408 |
configuration for a particular installation uses more specific names derived |
71131 | 409 |
from \<^verbatim>\<open>phabricator-\<close>\<open>NAME\<close>, e.g.\ \<^verbatim>\<open>/var/www/phabricator-vcs\<close> for the |
410 |
default. |
|
71103 | 411 |
|
71131 | 412 |
Knowing the naming conventions, it is possible to purge a Linux installation |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
413 |
from Isabelle/Phorge with some effort, but there is no automated |
71131 | 414 |
procedure for de-installation. In the worst case, it might be better to |
415 |
re-install the virtual machine from a clean image. |
|
71103 | 416 |
|
417 |
\<^medskip> |
|
418 |
Option \<^verbatim>\<open>-U\<close> ensures a full update of system packages, before installing |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
419 |
further packages required by Phorge. This might require a reboot. |
71103 | 420 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
421 |
Option \<^verbatim>\<open>-M:\<close> installs a standard Mercurial release from source: a specific |
81724 | 422 |
version that is known to work on Ubuntu 22.04 or 24.04, respectively. It is |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
423 |
also possible to specify the path or URL of the source archive (\<^verbatim>\<open>.tar.gz\<close>). |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
424 |
This option is recommended for production use, but it requires to |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
425 |
\<^emph>\<open>uninstall\<close> existing Mercurial packages provided by the operating system. |
71280
5a2033fc8f3d
avoid odd (harmless) problem with Mercurial 4.5.3 provided by Ubuntu 18.04 on first push: "couldn't write revision branch cache names";
wenzelm
parents:
71279
diff
changeset
|
426 |
|
71103 | 427 |
Option \<^verbatim>\<open>-n\<close> provides an alternative installation name. The default name |
71131 | 428 |
\<^verbatim>\<open>vcs\<close> means ``version control system''. The name appears in the URL for SSH |
429 |
access, and thus has some relevance to end-users. The initial server URL |
|
430 |
also uses the same suffix, but that can (and should) be changed later via |
|
431 |
regular Apache configuration. |
|
71103 | 432 |
|
71422
5d5be87330b5
allow to override repository versions at runtime;
wenzelm
parents:
71390
diff
changeset
|
433 |
Option \<^verbatim>\<open>-o\<close> augments the environment of Isabelle system options: relevant |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
434 |
options for Isabelle/Phorge have the prefix ``\<^verbatim>\<open>phabricator_\<close>'' (see |
71422
5d5be87330b5
allow to override repository versions at runtime;
wenzelm
parents:
71390
diff
changeset
|
435 |
also the result of e.g. ``\<^verbatim>\<open>isabelle options -l\<close>''). |
5d5be87330b5
allow to override repository versions at runtime;
wenzelm
parents:
71390
diff
changeset
|
436 |
|
71103 | 437 |
Option \<^verbatim>\<open>-r\<close> specifies an alternative installation root directory: it needs |
438 |
to be accessible for the Apache web server. |
|
439 |
||
440 |
Option \<^verbatim>\<open>-R\<close> specifies an alternative directory for repositories that are |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
441 |
hosted by Phorge. Provided that it is accessible for the Apache web |
71131 | 442 |
server, the directory can be reused for the \<^verbatim>\<open>hgweb\<close> view by Mercurial.\<^footnote>\<open>See |
443 |
also the documentation |
|
444 |
\<^url>\<open>https://www.mercurial-scm.org/wiki/PublishingRepositories\<close> and the |
|
445 |
example \<^url>\<open>https://isabelle.sketis.net/repos\<close>.\<close> |
|
71103 | 446 |
\<close> |
447 |
||
448 |
||
449 |
subsection \<open>\<^verbatim>\<open>isabelle phabricator_setup_mail\<close>\<close> |
|
450 |
||
451 |
text \<open> |
|
71290 | 452 |
The @{tool_def phabricator_setup_mail} tool provides mail configuration for |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
453 |
an existing Phorge installation: |
71103 | 454 |
@{verbatim [display] \<open>Usage: isabelle phabricator_setup_mail [OPTIONS] |
455 |
||
456 |
Options are: |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
457 |
-T USER send test mail to Phorge user |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
458 |
-f FILE config file (default: "mailers.json" within Phorge root) |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
459 |
-n NAME Phorge installation name (default: "vcs") |
71103 | 460 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
461 |
Provide mail configuration for existing Phorge installation.\<close>} |
71103 | 462 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
463 |
Proper mail configuration is vital for Phorge, but the details can be |
71103 | 464 |
tricky. A common approach is to re-use an existing SMTP mail service, as is |
71131 | 465 |
often included in regular web hosting packages. It is sufficient to create |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
466 |
one mail account for multiple Phorge installations, but the |
71131 | 467 |
configuration needs to be set for each installation. |
71103 | 468 |
|
469 |
The first invocation of \<^verbatim>\<open>isabelle phabricator_setup_mail\<close> without options |
|
470 |
creates a JSON template file. Its \<^verbatim>\<open>key\<close> entry should be changed to |
|
471 |
something sensible to identify the configuration, e.g.\ the Internet Domain |
|
71131 | 472 |
Name of the mail address. The \<^verbatim>\<open>options\<close> specify the SMTP server address and |
473 |
account information. |
|
71103 | 474 |
|
475 |
Another invocation of \<^verbatim>\<open>isabelle phabricator_setup_mail\<close> with updated JSON |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
476 |
file will change the underlying Phorge installation. This can be done |
71103 | 477 |
repeatedly, until everything works as expected. |
478 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
479 |
Option \<^verbatim>\<open>-T\<close> invokes a standard Phorge test procedure for the mail |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
480 |
configuration. The argument needs to be a valid Phorge user: the mail |
71103 | 481 |
address is derived from the user profile. |
482 |
||
483 |
Option \<^verbatim>\<open>-f\<close> refers to an existing JSON configuration file, e.g.\ from a |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
484 |
previous successful Phorge installation: sharing mailers setup with the |
71131 | 485 |
same mail address is fine for outgoing mails; incoming mails are optional |
486 |
and not configured here. |
|
71099 | 487 |
\<close> |
488 |
||
71132 | 489 |
|
490 |
subsection \<open>\<^verbatim>\<open>isabelle phabricator_setup_ssh\<close>\<close> |
|
491 |
||
492 |
text \<open> |
|
71290 | 493 |
The @{tool_def phabricator_setup_ssh} tool configures a special SSH service |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
494 |
for all Phorge installations: |
71132 | 495 |
@{verbatim [display] \<open>Usage: isabelle phabricator_setup_ssh [OPTIONS] |
496 |
||
497 |
Options are: |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
498 |
-p PORT sshd port for Phorge servers (default: 2222) |
71132 | 499 |
-q PORT sshd port for the operating system (default: 22) |
500 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
501 |
Configure ssh service for all Phorge installations: a separate sshd |
71132 | 502 |
is run in addition to the one of the operating system, and ports need to |
503 |
be distinct. |
|
504 |
||
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
505 |
A particular Phorge installation is addressed by using its |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
506 |
name as the ssh user; the actual Phorge user is determined via |
71132 | 507 |
stored ssh keys.\<close>} |
508 |
||
509 |
This is optional, but very useful. It allows to refer to hosted repositories |
|
510 |
via ssh with the usual public-key authentication. It also allows to |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
511 |
communicate with a Phorge server via the JSON API of |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
512 |
\<^emph>\<open>Conduit\<close>\<^footnote>\<open>\<^url>\<open>https://we.phorge.it/book/phorge/article/conduit\<close>\<close>. |
71132 | 513 |
|
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
514 |
\<^medskip> The Phorge SSH server distinguishes installations by their name, |
71132 | 515 |
e.g.\ \<^verbatim>\<open>vcs\<close> as SSH user name. The public key that is used for |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
516 |
authentication identifies the user within Phorge: there is a web |
71132 | 517 |
interface to provide that as part of the user profile. |
518 |
||
519 |
The operating system already has an SSH server (by default on port 22) that |
|
520 |
remains important for remote administration of the machine. |
|
521 |
||
522 |
\<^medskip> |
|
71292
8b745b4d71b5
more rebust setup for Subversion -- assuming that diffusion.ssh-port is only set in phabricator/conf/local/local.json, not in the database;
wenzelm
parents:
71290
diff
changeset
|
523 |
Options \<^verbatim>\<open>-p\<close> and \<^verbatim>\<open>-q\<close> allow to change the port assignment for both |
8b745b4d71b5
more rebust setup for Subversion -- assuming that diffusion.ssh-port is only set in phabricator/conf/local/local.json, not in the database;
wenzelm
parents:
71290
diff
changeset
|
524 |
servers. A common scheme is \<^verbatim>\<open>-p 22 -q 222\<close> to leave the standard port to |
81717
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
525 |
Phorge, to simplify the ssh URL that users will see for remote repository |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
526 |
clones.\<^footnote>\<open>For the rare case of hosting Subversion repositories, port 22 is |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
527 |
de-facto required. Otherwise Phorge presents malformed \<^verbatim>\<open>svn+ssh\<close> URLs with |
114449035ec6
update NEWS / documentation / descriptions for Phorge (formerly Phabricator);
wenzelm
parents:
79496
diff
changeset
|
528 |
port specification.\<close> |
71132 | 529 |
|
530 |
Redirecting the operating system sshd to port 222 requires some care: it |
|
531 |
requires to adjust the remote login procedure, e.g.\ in \<^verbatim>\<open>$HOME/.ssh/config\<close> |
|
532 |
to add a \<^verbatim>\<open>Port\<close> specification for the server machine. |
|
533 |
\<close> |
|
534 |
||
71099 | 535 |
end |